diff --git a/thys/Containers/ITP-2013/Benchmark_Set_LC.thy b/thys/Containers/ITP-2013/Benchmark_Set_LC.thy --- a/thys/Containers/ITP-2013/Benchmark_Set_LC.thy +++ b/thys/Containers/ITP-2013/Benchmark_Set_LC.thy @@ -1,88 +1,88 @@ theory Benchmark_Set_LC imports Benchmark_Set Containers.Set_Impl "HOL-Library.Code_Target_Nat" begin lemma [code_unfold del]: "card \ Cardinality.card'" by(simp) instantiation word :: (len) ceq begin definition "CEQ('a word) = Some (=)" instance by(intro_classes)(simp add: ceq_word_def) end instantiation word :: (len) compare begin definition "compare_word = (comparator_of :: 'a word comparator)" instance by(intro_classes)(simp add: compare_word_def comparator_of) end instantiation word :: (len) ccompare begin definition "CCOMPARE('a word) = Some compare" instance by(intro_classes)(simp add: ccompare_word_def comparator_compare) end instantiation word :: (len) set_impl begin definition "SET_IMPL('a word) = Phantom('a word) set_RBT" instance .. end instantiation word :: (len) proper_interval begin fun proper_interval_word :: "'a word option \ 'a word option \ bool" where "proper_interval_word None None = True" | "proper_interval_word None (Some y) = (y \ 0)" -| "proper_interval_word (Some x) None = (x \ max_word)" +| "proper_interval_word (Some x) None = (x \ - 1)" | "proper_interval_word (Some x) (Some y) = (x < y \ x \ y - 1)" instance proof let ?pi = "proper_interval :: 'a word proper_interval" show "?pi None None = True" by simp fix y show "?pi None (Some y) = (\z. z < y)" - by simp (metis word_gt_0 word_not_simps(1)) + using word_neq_0_conv [of y] by auto fix x show "?pi (Some x) None = (\z. x < z)" - by simp (metis eq_iff max_word_max not_le) + using word_order.not_eq_extremum [of x] by auto have "(x < y \ x \ y - 1) = (\z>x. z < y)" (is "?lhs \ ?rhs") proof assume ?lhs then obtain "x < y" "x \ y - 1" .. have "0 \ uint x" by simp also have "\ < uint y" using \x < y\ by(simp add: word_less_def) finally have "0 < uint y" . then have "y - 1 < y" by(simp add: word_less_def uint_sub_if' not_le) moreover from \0 < uint y\ \x < y\ \x \ y - 1\ have "x < y - 1" by(simp add: word_less_def uint_sub_if' uint_arith_simps(3)) ultimately show ?rhs by blast next assume ?rhs then obtain z where z: "x < z" "z < y" by blast have "0 \ uint z" by simp also have "\ < uint y" using \z < y\ by(simp add: word_less_def) finally show ?lhs using z by(auto simp add: word_less_def uint_sub_if') qed thus "?pi (Some x) (Some y) = (\z>x. z < y)" by simp qed end instantiation word :: (len) cproper_interval begin definition "cproper_interval = (proper_interval :: 'a word proper_interval)" instance by( intro_classes, simp add: cproper_interval_word_def ccompare_word_def compare_word_def le_lt_comparator_of ID_Some proper_interval_class.axioms) end instantiation word :: (len) cenum begin definition "CENUM('a word) = None" instance by(intro_classes)(simp_all add: cEnum_word_def) end notepad begin have "Benchmark_Set.complete 30 40 (12345, 67899) = (32, 4294967296)" by eval end end diff --git a/thys/IEEE_Floating_Point/IEEE.thy b/thys/IEEE_Floating_Point/IEEE.thy --- a/thys/IEEE_Floating_Point/IEEE.thy +++ b/thys/IEEE_Floating_Point/IEEE.thy @@ -1,448 +1,448 @@ (* Formalization of IEEE-754 Standard for binary floating-point arithmetic *) (* Author: Lei Yu, University of Cambridge *) section \Specification of the IEEE standard\ theory IEEE imports "HOL-Library.Float" Word_Lib.Word_Lemmas begin typedef (overloaded) ('e::len, 'f::len) float = "UNIV::(1 word \ 'e word \ 'f word) set" by auto setup_lifting type_definition_float syntax "_float" :: "type \ type \ type" ("'(_, _') float") text \parse \('a, 'b) float\ as ('a::len, 'b::len) float.\ parse_translation \ let fun float t u = Syntax.const @{type_syntax float} $ t $ u; fun len_tr u = (case Term_Position.strip_positions u of v as Free (x, _) => if Lexicon.is_tid x then (Syntax.const @{syntax_const "_ofsort"} $ v $ Syntax.const @{class_syntax len}) else u | _ => u) fun len_float_tr [t, u] = float (len_tr t) (len_tr u) in [(@{syntax_const "_float"}, K len_float_tr)] end \ subsection \Derived parameters for floating point formats\ definition wordlength :: "('e, 'f) float itself \ nat" where "wordlength x = LENGTH('e) + LENGTH('f) + 1" definition bias :: "('e, 'f) float itself \ nat" where "bias x = 2^(LENGTH('e) - 1) - 1" definition emax :: "('e, 'f) float itself \ nat" - where "emax x = unat(max_word::'e word)" + where "emax x = unat (- 1::'e word)" abbreviation fracwidth::"('e, 'f) float itself \ nat" where "fracwidth _ \ LENGTH('f)" subsection \Predicates for the four IEEE formats\ definition is_single :: "('e, 'f) float itself \ bool" where "is_single x \ LENGTH('e) = 8 \ wordlength x = 32" definition is_double :: "('e, 'f) float itself \ bool" where "is_double x \ LENGTH('e) = 11 \ wordlength x = 64" definition is_single_extended :: "('e, 'f) float itself \ bool" where "is_single_extended x \ LENGTH('e) \ 11 \ wordlength x \ 43" definition is_double_extended :: "('e, 'f) float itself \ bool" where "is_double_extended x \ LENGTH('e) \ 15 \ wordlength x \ 79" subsection \Extractors for fields\ lift_definition sign::"('e, 'f) float \ nat" is "\(s::1 word, _::'e word, _::'f word). unat s" . lift_definition exponent::"('e, 'f) float \ nat" is "\(_, e::'e word, _). unat e" . lift_definition fraction::"('e, 'f) float \ nat" is "\(_, _, f::'f word). unat f" . abbreviation "real_of_word x \ real (unat x)" lift_definition valof :: "('e, 'f) float \ real" is "\(s, e, f). let x = (TYPE(('e, 'f) float)) in (if e = 0 then (-1::real)^(unat s) * (2 / (2^bias x)) * (real_of_word f/2^(LENGTH('f))) else (-1::real)^(unat s) * ((2^(unat e)) / (2^bias x)) * (1 + real_of_word f/2^LENGTH('f)))" . subsection \Partition of numbers into disjoint classes\ definition is_nan :: "('e, 'f) float \ bool" where "is_nan a \ exponent a = emax TYPE(('e, 'f)float) \ fraction a \ 0" definition is_infinity :: "('e, 'f) float \ bool" where "is_infinity a \ exponent a = emax TYPE(('e, 'f)float) \ fraction a = 0" definition is_normal :: "('e, 'f) float \ bool" where "is_normal a \ 0 < exponent a \ exponent a < emax TYPE(('e, 'f)float)" definition is_denormal :: "('e, 'f) float \ bool" where "is_denormal a \ exponent a = 0 \ fraction a \ 0" definition is_zero :: "('e, 'f) float \ bool" where "is_zero a \ exponent a = 0 \ fraction a = 0" definition is_finite :: "('e, 'f) float \ bool" where "is_finite a \ (is_normal a \ is_denormal a \ is_zero a)" subsection \Special values\ -lift_definition plus_infinity :: "('e, 'f) float" ("\") is "(0, max_word, 0)" . +lift_definition plus_infinity :: "('e, 'f) float" ("\") is "(0, - 1, 0)" . -lift_definition topfloat :: "('e, 'f) float" is "(0, max_word - 1, 2^LENGTH('f) - 1)" . +lift_definition topfloat :: "('e, 'f) float" is "(0, - 2, 2^LENGTH('f) - 1)" . instantiation float::(len, len) zero begin lift_definition zero_float :: "('e, 'f) float" is "(0, 0, 0)" . instance proof qed end subsection \Negation operation on floating point values\ instantiation float::(len, len) uminus begin lift_definition uminus_float :: "('e, 'f) float \ ('e, 'f) float" is "\(s, e, f). (1 - s, e, f)" . instance proof qed end abbreviation (input) "minus_zero \ - (0::('e, 'f)float)" abbreviation (input) "minus_infinity \ - \" abbreviation (input) "bottomfloat \ - topfloat" subsection \Real number valuations\ text \The largest value that can be represented in floating point format.\ definition largest :: "('e, 'f) float itself \ real" where "largest x = (2^(emax x - 1) / 2^bias x) * (2 - 1/(2^fracwidth x))" text \Threshold, used for checking overflow.\ definition threshold :: "('e, 'f) float itself \ real" where "threshold x = (2^(emax x - 1) / 2^bias x) * (2 - 1/(2^(Suc(fracwidth x))))" text \Unit of least precision.\ lift_definition one_lp::"('e ,'f) float \ ('e ,'f) float" is "\(s, e, f). (0, e::'e word, 1)" . lift_definition zero_lp::"('e ,'f) float \ ('e ,'f) float" is "\(s, e, f). (0, e::'e word, 0)" . definition ulp :: "('e, 'f) float \ real" where "ulp a = valof (one_lp a) - valof (zero_lp a)" text \Enumerated type for rounding modes.\ datatype roundmode = To_nearest | float_To_zero | To_pinfinity | To_ninfinity text \Characterization of best approximation from a set of abstract values.\ definition "is_closest v s x a \ a \ s \ (\b. b \ s \ \v a - x\ \ \v b - x\)" text \Best approximation with a deciding preference for multiple possibilities.\ definition "closest v p s x = (SOME a. is_closest v s x a \ ((\b. is_closest v s x b \ p b) \ p a))" subsection \Rounding\ fun round :: "roundmode \ real \ ('e ,'f) float" where "round To_nearest y = (if y \ - threshold TYPE(('e ,'f) float) then minus_infinity else if y \ threshold TYPE(('e ,'f) float) then plus_infinity else closest (valof) (\a. even (fraction a)) {a. is_finite a} y)" | "round float_To_zero y = (if y < - largest TYPE(('e ,'f) float) then bottomfloat else if y > largest TYPE(('e ,'f) float) then topfloat else closest (valof) (\a. True) {a. is_finite a \ \valof a\ \ \y\} y)" | "round To_pinfinity y = (if y < - largest TYPE(('e ,'f) float) then bottomfloat else if y > largest TYPE(('e ,'f) float) then plus_infinity else closest (valof) (\a. True) {a. is_finite a \ valof a \ y} y)" | "round To_ninfinity y = (if y < - largest TYPE(('e ,'f) float) then minus_infinity else if y > largest TYPE(('e ,'f) float) then topfloat else closest (valof) (\a. True) {a. is_finite a \ valof a \ y} y)" text \Rounding to integer values in floating point format.\ definition is_integral :: "('e ,'f) float \ bool" where "is_integral a \ is_finite a \ (\n::nat. \valof a\ = real n)" fun intround :: "roundmode \ real \ ('e ,'f) float" where "intround To_nearest y = (if y \ - threshold TYPE(('e ,'f) float) then minus_infinity else if y \ threshold TYPE(('e ,'f) float) then plus_infinity else closest (valof) (\a. (\n::nat. even n \ \valof a\ = real n)) {a. is_integral a} y)" |"intround float_To_zero y = (if y < - largest TYPE(('e ,'f) float) then bottomfloat else if y > largest TYPE(('e ,'f) float) then topfloat else closest (valof) (\x. True) {a. is_integral a \ \valof a\ \ \y\} y)" |"intround To_pinfinity y = (if y < - largest TYPE(('e ,'f) float) then bottomfloat else if y > largest TYPE(('e ,'f) float) then plus_infinity else closest (valof) (\x. True) {a. is_integral a \ valof a \ y} y)" |"intround To_ninfinity y = (if y < - largest TYPE(('e ,'f) float) then minus_infinity else if y > largest TYPE(('e ,'f) float) then topfloat else closest (valof) (\x. True) {a. is_integral a \ valof a \ y} y)" text \Round, choosing between -0.0 or +0.0\ definition float_round::"roundmode \ bool \ real \ ('e, 'f) float" where "float_round mode toneg r = (let x = round mode r in if is_zero x then if toneg then minus_zero else 0 else x)" text \Non-standard of NaN.\ definition some_nan :: "('e ,'f) float" where "some_nan = (SOME a. is_nan a)" text \Coercion for signs of zero results.\ definition zerosign :: "nat \ ('e ,'f) float \ ('e ,'f) float" where "zerosign s a = (if is_zero a then (if s = 0 then 0 else - 0) else a)" text \Remainder operation.\ definition rem :: "real \ real \ real" where "rem x y = (let n = closest id (\x. \n::nat. even n \ \x\ = real n) {x. \n :: nat. \x\ = real n} (x / y) in x - n * y)" definition frem :: "roundmode \ ('e ,'f) float \ ('e ,'f) float \ ('e ,'f) float" where "frem m a b = (if is_nan a \ is_nan b \ is_infinity a \ is_zero b then some_nan else zerosign (sign a) (round m (rem (valof a) (valof b))))" subsection \Definitions of the arithmetic operations\ definition fintrnd :: "roundmode \ ('e ,'f) float \ ('e ,'f) float" where "fintrnd m a = (if is_nan a then (some_nan) else if is_infinity a then a else zerosign (sign a) (intround m (valof a)))" definition fadd :: "roundmode \ ('e ,'f) float \ ('e ,'f) float \ ('e ,'f) float" where "fadd m a b = (if is_nan a \ is_nan b \ (is_infinity a \ is_infinity b \ sign a \ sign b) then some_nan else if (is_infinity a) then a else if (is_infinity b) then b else zerosign (if is_zero a \ is_zero b \ sign a = sign b then sign a else if m = To_ninfinity then 1 else 0) (round m (valof a + valof b)))" definition fsub :: "roundmode \ ('e ,'f) float \ ('e ,'f) float \ ('e ,'f) float" where "fsub m a b = (if is_nan a \ is_nan b \ (is_infinity a \ is_infinity b \ sign a = sign b) then some_nan else if is_infinity a then a else if is_infinity b then - b else zerosign (if is_zero a \ is_zero b \ sign a \ sign b then sign a else if m = To_ninfinity then 1 else 0) (round m (valof a - valof b)))" definition fmul :: "roundmode \ ('e ,'f) float \ ('e ,'f) float \ ('e ,'f) float" where "fmul m a b = (if is_nan a \ is_nan b \ (is_zero a \ is_infinity b) \ (is_infinity a \ is_zero b) then some_nan else if is_infinity a \ is_infinity b then (if sign a = sign b then plus_infinity else minus_infinity) else zerosign (if sign a = sign b then 0 else 1 ) (round m (valof a * valof b)))" definition fdiv :: "roundmode \ ('e ,'f) float \ ('e ,'f) float \ ('e ,'f) float" where "fdiv m a b = (if is_nan a \ is_nan b \ (is_zero a \ is_zero b) \ (is_infinity a \ is_infinity b) then some_nan else if is_infinity a \ is_zero b then (if sign a = sign b then plus_infinity else minus_infinity) else if is_infinity b then (if sign a = sign b then 0 else - 0) else zerosign (if sign a = sign b then 0 else 1) (round m (valof a / valof b)))" definition fsqrt :: "roundmode \ ('e ,'f) float \ ('e ,'f) float" where "fsqrt m a = (if is_nan a then some_nan else if is_zero a \ is_infinity a \ sign a = 0 then a else if sign a = 1 then some_nan else zerosign (sign a) (round m (sqrt (valof a))))" definition fmul_add :: "roundmode \ ('t ,'w) float \ ('t ,'w) float \ ('t ,'w) float \ ('t ,'w) float" where "fmul_add mode x y z = (let signP = if sign x = sign y then 0 else 1 in let infP = is_infinity x \ is_infinity y in if is_nan x \ is_nan y \ is_nan z then some_nan else if is_infinity x \ is_zero y \ is_zero x \ is_infinity y \ is_infinity z \ infP \ signP \ sign z then some_nan else if is_infinity z \ (sign z = 0) \ infP \ (signP = 0) then plus_infinity else if is_infinity z \ (sign z = 1) \ infP \ (signP = 1) then minus_infinity else let r1 = valof x * valof y; r2 = valof z in float_round mode (if (r1 = 0) \ (r2 = 0) \ (signP = sign z) then signP = 1 else mode = To_ninfinity) (r1 + r2))" subsection \Comparison operations\ datatype ccode = Gt | Lt | Eq | Und definition fcompare :: "('e ,'f) float \ ('e ,'f) float \ ccode" where "fcompare a b = (if is_nan a \ is_nan b then Und else if is_infinity a \ sign a = 1 then (if is_infinity b \ sign b = 1 then Eq else Lt) else if is_infinity a \ sign a = 0 then (if is_infinity b \ sign b = 0 then Eq else Gt) else if is_infinity b \ sign b = 1 then Gt else if is_infinity b \ sign b = 0 then Lt else if valof a < valof b then Lt else if valof a = valof b then Eq else Gt)" definition flt :: "('e ,'f) float \ ('e ,'f) float \ bool" where "flt a b \ fcompare a b = Lt" definition fle :: "('e ,'f) float \ ('e ,'f) float \ bool" where "fle a b \ fcompare a b = Lt \ fcompare a b = Eq" definition fgt :: "('e ,'f) float \ ('e ,'f) float \ bool" where "fgt a b \ fcompare a b = Gt" definition fge :: "('e ,'f) float \ ('e ,'f) float \ bool" where "fge a b \ fcompare a b = Gt \ fcompare a b = Eq" definition feq :: "('e ,'f) float \ ('e ,'f) float \ bool" where "feq a b \ fcompare a b = Eq" section \Specify float to be double precision and round to even\ instantiation float :: (len, len) plus begin definition plus_float :: "('a, 'b) float \ ('a, 'b) float \ ('a, 'b) float" where "a + b = fadd To_nearest a b" instance .. end instantiation float :: (len, len) minus begin definition minus_float :: "('a, 'b) float \ ('a, 'b) float \ ('a, 'b) float" where "a - b = fsub To_nearest a b" instance .. end instantiation float :: (len, len) times begin definition times_float :: "('a, 'b) float \ ('a, 'b) float \ ('a, 'b) float" where "a * b = fmul To_nearest a b" instance .. end instantiation float :: (len, len) one begin lift_definition one_float :: "('a, 'b) float" is "(0, 2^(LENGTH('a) - 1) - 1, 0)" . instance .. end instantiation float :: (len, len) inverse begin definition divide_float :: "('a, 'b) float \ ('a, 'b) float \ ('a, 'b) float" where "a div b = fdiv To_nearest a b" definition inverse_float :: "('a, 'b) float \ ('a, 'b) float" where "inverse_float a = fdiv To_nearest 1 a" instance .. end definition float_rem :: "('a, 'b) float \ ('a, 'b) float \ ('a, 'b) float" where "float_rem a b = frem To_nearest a b" definition float_sqrt :: "('a, 'b) float \ ('a, 'b) float" where "float_sqrt a = fsqrt To_nearest a" definition ROUNDFLOAT ::"('a, 'b) float \ ('a, 'b) float" where "ROUNDFLOAT a = fintrnd To_nearest a" instantiation float :: (len, len) ord begin definition less_float :: "('a, 'b) float \ ('a, 'b) float \ bool" where "a < b \ flt a b" definition less_eq_float :: "('a, 'b) float \ ('a, 'b) float \ bool" where "a \ b \ fle a b" instance .. end definition float_eq :: "('a, 'b) float \ ('a, 'b) float \ bool" (infixl "\" 70) where "float_eq a b = feq a b" instantiation float :: (len, len) abs begin definition abs_float :: "('a, 'b) float \ ('a, 'b) float" where "abs_float a = (if sign a = 0 then a else - a)" instance .. end text \The \1 + \\ property.\ definition normalizes :: "_ itself \ real \ bool" where "normalizes float_format x = (1/ (2::real)^(bias float_format - 1) \ \x\ \ \x\ < threshold float_format)" end diff --git a/thys/IEEE_Floating_Point/IEEE_Properties.thy b/thys/IEEE_Floating_Point/IEEE_Properties.thy --- a/thys/IEEE_Floating_Point/IEEE_Properties.thy +++ b/thys/IEEE_Floating_Point/IEEE_Properties.thy @@ -1,1023 +1,1023 @@ (* Author: Lei Yu, University of Cambridge Author: Fabian Hellauer Fabian Immler *) section \Proofs of Properties about Floating Point Arithmetic\ theory IEEE_Properties imports IEEE begin subsection \Theorems derived from definitions\ lemma valof_eq: "valof x = (if exponent x = 0 then (- 1) ^ sign x * (2 / 2 ^ bias TYPE(('a, 'b) float)) * (real (fraction x) / 2 ^ LENGTH('b)) else (- 1) ^ sign x * (2 ^ exponent x / 2 ^ bias TYPE(('a, 'b) float)) * (1 + real (fraction x) / 2 ^ LENGTH('b)))" for x::"('a, 'b) float" unfolding Let_def by transfer (auto simp: bias_def divide_simps unat_eq_0) lemma exponent_le[simp]: - "exponent a \ unat (max_word::'a word)" for a::"('a, _)float" + "exponent a \ unat (-1::'a word)" for a::"('a, _)float" by transfer (auto intro!: unat_mono simp: word_le_nat_alt[symmetric]) lemma max_word_le_exponent_iff[simp]: - "unat (max_word::'a word) \ exponent a \ unat (max_word::'a word) = exponent a" + "unat (- 1::'a word) \ exponent a \ unat (- 1::'a word) = exponent a" for a::"('a, _)float" using le_antisym by fastforce lemma infinity_simps: "sign (plus_infinity::('e, 'f)float) = 0" "sign (minus_infinity::('e, 'f)float) = 1" "exponent (plus_infinity::('e, 'f)float) = emax TYPE(('e, 'f)float)" "exponent (minus_infinity::('e, 'f)float) = emax TYPE(('e, 'f)float)" "fraction (plus_infinity::('e, 'f)float) = 0" "fraction (minus_infinity::('e, 'f)float) = 0" subgoal by transfer auto subgoal by transfer auto subgoal by transfer (auto simp: emax_def) subgoal by transfer (auto simp: emax_def) subgoal by transfer auto subgoal by transfer auto done lemma zero_simps: "sign (0::('e, 'f)float) = 0" "sign (- 0::('e, 'f)float) = 1" "exponent (0::('e, 'f)float) = 0" "exponent (- 0::('e, 'f)float) = 0" "fraction (0::('e, 'f)float) = 0" "fraction (- 0::('e, 'f)float) = 0" subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto done lemma emax_eq: "emax x = 2^LENGTH('e) - 1" for x::"('e, 'f)float itself" by (simp add: emax_def unat_minus_one_word) lemma topfloat_simps: "sign (topfloat::('e, 'f)float) = 0" "exponent (topfloat::('e, 'f)float) = emax TYPE(('e, 'f)float) - 1" "fraction (topfloat::('e, 'f)float) = 2^fracwidth TYPE(('e, 'f)float) - 1" and bottomfloat_simps: "sign (bottomfloat::('e, 'f)float) = 1" "exponent (bottomfloat::('e, 'f)float) = emax TYPE(('e, 'f)float) - 1" "fraction (bottomfloat::('e, 'f)float) = 2^fracwidth TYPE(('e, 'f)float) - 1" subgoal by transfer simp subgoal by transfer (simp add: emax_eq take_bit_minus_small_eq nat_diff_distrib nat_power_eq) subgoal by transfer (simp add: unat_minus_one_word) subgoal by transfer simp subgoal by transfer (simp add: emax_eq take_bit_minus_small_eq unat_minus_one_word nat_diff_distrib nat_power_eq) subgoal by transfer (simp add: unat_minus_one_word) done lemmas float_defs = is_finite_def is_infinity_def is_zero_def is_nan_def is_normal_def is_denormal_def valof_eq less_eq_float_def less_float_def flt_def fgt_def fle_def fge_def feq_def fcompare_def infinity_simps zero_simps topfloat_simps bottomfloat_simps float_eq_def lemma float_cases: "is_nan a \ is_infinity a \ is_normal a \ is_denormal a \ is_zero a" by (auto simp: emax_def float_defs not_less) lemma float_cases_finite: "is_nan a \ is_infinity a \ is_finite a" by (simp add: float_cases is_finite_def) lemma float_zero1[simp]: "is_zero 0" unfolding float_defs by transfer auto lemma float_zero2[simp]: "is_zero (- x) \ is_zero x" unfolding float_defs by transfer auto lemma emax_pos[simp]: "0 < emax x" "emax x \ 0" by (auto simp: emax_def) text \The types of floating-point numbers are mutually distinct.\ lemma float_distinct: "\ (is_nan a \ is_infinity a)" "\ (is_nan a \ is_normal a)" "\ (is_nan a \ is_denormal a)" "\ (is_nan a \ is_zero a)" "\ (is_infinity a \ is_normal a)" "\ (is_infinity a \ is_denormal a)" "\ (is_infinity a \ is_zero a)" "\ (is_normal a \ is_denormal a)" "\ (is_denormal a \ is_zero a)" by (auto simp: float_defs) lemma denormal_imp_not_zero: "is_denormal f \ \is_zero f" by (simp add: is_denormal_def is_zero_def) lemma normal_imp_not_zero: "is_normal f \ \is_zero f" by (simp add: is_normal_def is_zero_def) lemma normal_imp_not_denormal: "is_normal f \ \is_denormal f" by (simp add: is_normal_def is_denormal_def) lemma denormal_zero[simp]: "\is_denormal 0" "\is_denormal minus_zero" using denormal_imp_not_zero float_zero1 float_zero2 by blast+ lemma normal_zero[simp]: "\is_normal 0" "\is_normal minus_zero" using normal_imp_not_zero float_zero1 float_zero2 by blast+ lemma float_distinct_finite: "\ (is_nan a \ is_finite a)" "\(is_infinity a \ is_finite a)" by (auto simp: float_defs) lemma finite_infinity: "is_finite a \ \ is_infinity a" by (auto simp: float_defs) lemma finite_nan: "is_finite a \ \ is_nan a" by (auto simp: float_defs) text \For every real number, the floating-point numbers closest to it always exist.\ lemma is_closest_exists: fixes v :: "('e, 'f)float \ real" and s :: "('e, 'f)float set" assumes finite: "finite s" and non_empty: "s \ {}" shows "\a. is_closest v s x a" using finite non_empty proof (induct s rule: finite_induct) case empty then show ?case by simp next case (insert z s) show ?case proof (cases "s = {}") case True then have "is_closest v (insert z s) x z" by (auto simp: is_closest_def) then show ?thesis by metis next case False then obtain a where a: "is_closest v s x a" using insert by metis then show ?thesis proof (cases "\v a - x\" "\v z - x\" rule: le_cases) case le then show ?thesis by (metis insert_iff a is_closest_def) next case ge have "\b. b \ s \ \v a - x\ \ \v b - x\" by (metis a is_closest_def) then have "\b. b \ insert z s \ \v z - x\ \ \v b - x\" by (metis eq_iff ge insert_iff order.trans) then show ?thesis using is_closest_def a by (metis insertI1) qed qed qed lemma closest_is_everything: fixes v :: "('e, 'f)float \ real" and s :: "('e, 'f)float set" assumes finite: "finite s" and non_empty: "s \ {}" shows "is_closest v s x (closest v p s x) \ ((\b. is_closest v s x b \ p b) \ p (closest v p s x))" unfolding closest_def by (rule someI_ex) (metis assms is_closest_exists [of s v x]) lemma closest_in_set: fixes v :: "('e, 'f)float \ real" assumes "finite s" and "s \ {}" shows "closest v p s x \ s" by (metis assms closest_is_everything is_closest_def) lemma closest_is_closest_finite: fixes v :: "('e, 'f)float \ real" assumes "finite s" and "s \ {}" shows "is_closest v s x (closest v p s x)" by (metis closest_is_everything assms) instance float::(len, len) finite proof qed (transfer, simp) lemma is_finite_nonempty: "{a. is_finite a} \ {}" proof - have "0 \ {a. is_finite a}" unfolding float_defs by transfer auto then show ?thesis by (metis empty_iff) qed lemma closest_is_closest: fixes v :: "('e, 'f)float \ real" assumes "s \ {}" shows "is_closest v s x (closest v p s x)" by (rule closest_is_closest_finite) (auto simp: assms) subsection \Properties about ordering and bounding\ text \Lifting of non-exceptional comparisons.\ lemma float_lt [simp]: assumes "is_finite a" "is_finite b" shows "a < b \ valof a < valof b" proof assume "valof a < valof b" moreover have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) ultimately have "fcompare a b = Lt" by (auto simp add: is_infinity_def is_nan_def valof_def fcompare_def) then show "a < b" by (auto simp: float_defs) next assume "a < b" then have lt: "fcompare a b = Lt" by (simp add: float_defs) have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then show "valof a < valof b" using lt assms by (simp add: fcompare_def is_nan_def is_infinity_def valof_def split: if_split_asm) qed lemma float_eq [simp]: assumes "is_finite a" "is_finite b" shows "a \ b \ valof a = valof b" proof assume *: "valof a = valof b" have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto with * have "fcompare a b = Eq" by (auto simp add: is_infinity_def is_nan_def valof_def fcompare_def) then show "a \ b" by (auto simp: float_defs) next assume "a \ b" then have eq: "fcompare a b = Eq" by (simp add: float_defs) have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto then show "valof a = valof b" using eq assms by (simp add: fcompare_def is_nan_def is_infinity_def valof_def split: if_split_asm) qed lemma float_le [simp]: assumes "is_finite a" "is_finite b" shows "a \ b \ valof a \ valof b" proof - have "a \ b \ a < b \ a \ b" by (auto simp add: float_defs) then show ?thesis by (auto simp add: assms) qed text \Reflexivity of equality for non-NaNs.\ lemma float_eq_refl [simp]: "a \ a \ \ is_nan a" by (auto simp: float_defs) text \Properties about Ordering.\ lemma float_lt_trans: "is_finite a \ is_finite b \ is_finite c \ a < b \ b < c \ a < c" by (auto simp: le_trans) lemma float_le_less_trans: "is_finite a \ is_finite b \ is_finite c \ a \ b \ b < c \ a < c" by (auto simp: le_trans) lemma float_le_trans: "is_finite a \ is_finite b \ is_finite c \ a \ b \ b \ c \ a \ c" by (auto simp: le_trans) lemma float_le_neg: "is_finite a \ is_finite b \ \ a < b \ b \ a" by auto text \Properties about bounding.\ lemma float_le_infinity [simp]: "\ is_nan a \ a \ plus_infinity" unfolding float_defs by auto lemma zero_le_topfloat[simp]: "0 \ topfloat" "- 0 \ topfloat" by (auto simp: float_defs field_simps power_gt1_lemma of_nat_diff) lemma LENGTH_contr: "Suc 0 < LENGTH('e) \ 2 ^ LENGTH('e::len) \ Suc (Suc 0) \ False" by (metis le_antisym len_gt_0 n_less_equal_power_2 not_less_eq numeral_2_eq_2 one_le_numeral self_le_power) lemma valof_topfloat: "valof (topfloat::('e, 'f)float) = largest TYPE(('e, 'f)float)" if "LENGTH('e) > 1" using that LENGTH_contr by (auto simp add: emax_eq largest_def divide_simps float_defs of_nat_diff) lemma float_frac_le: "fraction a \ 2^LENGTH('f) - 1" for a::"('e, 'f)float" unfolding float_defs using less_Suc_eq_le by transfer fastforce lemma float_exp_le: "is_finite a \ exponent a \ emax TYPE(('e, 'f)float) - 1" for a::"('e, 'f)float" unfolding float_defs by auto lemma float_sign_le: "(-1::real)^(sign a) = 1 \ (-1::real)^(sign a) = -1" by (metis neg_one_even_power neg_one_odd_power) lemma exp_less: "a \ b \ (2::real)^a \ 2^b" for a b :: nat by auto lemma div_less: "a \ b \ c > 0 \ a/c \ b/c" for a b c :: "'a::linordered_field" by (metis divide_le_cancel less_asym) lemma finite_topfloat: "is_finite topfloat" unfolding float_defs by auto lemmas float_leI = float_le[THEN iffD2] lemma factor_minus: "x * a - x = x * (a - 1)" for x a::"'a::comm_semiring_1_cancel" by (simp add: algebra_simps) lemma real_le_power_numeral_diff: "real a \ numeral b ^ n - 1 \ a \ numeral b ^ n - 1" by (metis (mono_tags, lifting) of_nat_1 of_nat_diff of_nat_le_iff of_nat_numeral one_le_numeral one_le_power semiring_1_class.of_nat_power) definition denormal_exponent::"('e, 'f)float itself \ int" where "denormal_exponent x = 1 - (int (LENGTH('f)) + int (bias TYPE(('e, 'f)float)))" definition normal_exponent::"('e, 'f)float \ int" where "normal_exponent x = int (exponent x) - int (bias TYPE(('e, 'f)float)) - int (LENGTH('f))" definition denormal_mantissa::"('e, 'f)float \ int" where "denormal_mantissa x = (-1::int)^sign x * int (fraction x)" definition normal_mantissa::"('e, 'f)float \ int" where "normal_mantissa x = (-1::int)^sign x * (2^LENGTH('f) + int (fraction x))" lemma unat_one_word_le: "unat a \ Suc 0" for a::"1 word" using unat_lt2p[of a] by auto lemma one_word_le: "a \ 1" for a::"1 word" by (auto simp: word_le_nat_alt unat_one_word_le) lemma sign_cases[case_names pos neg]: obtains "sign x = 0" | "sign x = 1" proof cases assume "sign x = 0" then show ?thesis .. next assume "sign x \ 0" have "sign x \ 1" by transfer (auto simp: unat_one_word_le) then have "sign x = 1" using \sign x \ 0\ by auto then show ?thesis .. qed lemma is_infinity_cases: assumes "is_infinity x" obtains "x = plus_infinity" | "x = minus_infinity" proof (cases rule: sign_cases[of x]) assume "sign x = 0" then have "x = plus_infinity" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_0) then show ?thesis .. next assume "sign x = 1" then have "x = minus_infinity" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_of_nat) then show ?thesis .. qed lemma is_zero_cases: assumes "is_zero x" obtains "x = 0" | "x = - 0" proof (cases rule: sign_cases[of x]) assume "sign x = 0" then have "x = 0" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_0) then show ?thesis .. next assume "sign x = 1" then have "x = minus_zero" using assms unfolding float_defs by transfer (auto simp: emax_def unat_eq_of_nat) then show ?thesis .. qed lemma minus_minus_float[simp]: "- (-f) = f" for f::"('e, 'f)float" by transfer auto lemma sign_minus_float: "sign (-f) = (1 - sign f)" for f::"('e, 'f)float" by transfer (auto simp: unat_eq_1 one_word_le unat_sub) lemma exponent_uminus[simp]: "exponent (- f) = exponent f" by transfer auto lemma fraction_uminus[simp]: "fraction (- f) = fraction f" by transfer auto lemma is_normal_minus_float[simp]: "is_normal (-f) = is_normal f" for f::"('e, 'f)float" by (auto simp: is_normal_def) lemma is_denormal_minus_float[simp]: "is_denormal (-f) = is_denormal f" for f::"('e, 'f)float" by (auto simp: is_denormal_def) lemma bitlen_normal_mantissa: "bitlen (abs (normal_mantissa x)) = Suc LENGTH('f)" for x::"('e, 'f)float" proof - have "fraction x < 2 ^ LENGTH('f)" using float_frac_le[of x] by (metis One_nat_def Suc_pred le_imp_less_Suc pos2 zero_less_power) moreover have "- int (fraction x) \ 2 ^ LENGTH('f)" using negative_zle_0 order_trans zero_le_numeral zero_le_power by blast ultimately show ?thesis by (cases x rule: sign_cases) (auto simp: bitlen_le_iff_power bitlen_ge_iff_power nat_add_distrib normal_mantissa_def intro!: antisym) qed lemma less_int_natI: "x < y" if "0 \ x" "nat x < nat y" using that by arith lemma normal_exponent_bounds_int: "2 - 2 ^ (LENGTH('e) - 1) - int LENGTH('f) \ normal_exponent x" "normal_exponent x \ 2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1" if "is_normal x" for x::"('e, 'f)float" using that unfolding normal_exponent_def is_normal_def emax_eq bias_def by (auto simp del: zless_nat_conj intro!: less_int_natI simp: of_nat_diff nat_add_distrib nat_mult_distrib nat_power_eq power_Suc[symmetric]) lemmas of_int_leI = of_int_le_iff[THEN iffD2] lemma normal_exponent_bounds_real: "2 - 2 ^ (LENGTH('e) - 1) - real LENGTH('f) \ normal_exponent x" "normal_exponent x \ 2 ^ (LENGTH('e) - 1) - real LENGTH('f) - 1" if "is_normal x" for x::"('e, 'f)float" subgoal by (rule order_trans[OF _ of_int_leI[OF normal_exponent_bounds_int(1)[OF that]]]) auto subgoal by (rule order_trans[OF of_int_leI[OF normal_exponent_bounds_int(2)[OF that]]]) auto done lemma float_eqI: "x = y" if "sign x = sign y" "fraction x = fraction y" "exponent x = exponent y" using that by transfer auto lemma float_induct[induct type:float, case_names normal denormal neg zero infinity nan]: fixes a::"('e, 'f)float" assumes normal: "\x. is_normal x \ valof x = normal_mantissa x * 2 powr normal_exponent x \ P x" assumes denormal: "\x. is_denormal x \ valof x = denormal_mantissa x * 2 powr denormal_exponent TYPE(('e, 'f)float) \ P x" assumes zero: "P 0" "P minus_zero" assumes infty: "P plus_infinity" "P minus_infinity" assumes nan: "\x. is_nan x \ P x" shows "P a" proof - from float_cases[of a] consider "is_nan a" | "is_infinity a" | "is_normal a" | "is_denormal a" | "is_zero a" by blast then show ?thesis proof cases case 1 then show ?thesis by (rule nan) next case 2 then consider "a = plus_infinity" | "a = minus_infinity" by (rule is_infinity_cases) then show ?thesis by cases (auto intro: infty) next case hyps: 3 from hyps have "valof a = normal_mantissa a * 2 powr normal_exponent a" by (cases a rule: sign_cases) (auto simp: valof_eq normal_mantissa_def normal_exponent_def is_normal_def powr_realpow[symmetric] powr_diff powr_add divide_simps algebra_simps) from hyps this show ?thesis by (rule normal) next case hyps: 4 from hyps have "valof a = denormal_mantissa a * 2 powr denormal_exponent TYPE(('e, 'f)float)" by (cases a rule: sign_cases) (auto simp: valof_eq denormal_mantissa_def denormal_exponent_def is_denormal_def powr_realpow[symmetric] powr_diff powr_add divide_simps algebra_simps) from hyps this show ?thesis by (rule denormal) next case 5 then consider "a = 0" | "a = minus_zero" by (rule is_zero_cases) then show ?thesis by cases (auto intro: zero) qed qed lemma infinite_infinity[simp]: "\ is_finite plus_infinity" "\ is_finite minus_infinity" by (auto simp: is_finite_def is_normal_def infinity_simps is_denormal_def is_zero_def) lemma nan_not_finite[simp]: "is_nan x \ \ is_finite x" using float_distinct_finite(1) by blast lemma valof_nonneg: "valof x \ 0" if "sign x = 0" for x::"('e, 'f)float" by (auto simp: valof_eq that) lemma valof_nonpos: "valof x \ 0" if "sign x = 1" for x::"('e, 'f)float" using that by (auto simp: valof_eq is_finite_def) lemma real_le_intI: "x \ y" if "floor x \ floor y" "x \ \" for x y::real using that(2,1) by (induction rule: Ints_induct) (auto elim!: Ints_induct simp: le_floor_iff) lemma real_of_int_le_2_powr_bitlenI: "real_of_int x \ 2 powr n - 1" if "bitlen (abs x) \ m" "m \ n" proof - have "real_of_int x \ abs (real_of_int x)" by simp also have "\ < 2 powr (bitlen (abs x))" by (rule abs_real_le_2_powr_bitlen) finally have "real_of_int x \ 2 powr (bitlen (abs x)) - 1" by (auto simp: powr_real_of_int bitlen_nonneg intro!: real_le_intI) also have "\ \ 2 powr m - 1" by (simp add: that) also have "\ \ 2 powr n - 1" by (simp add: that) finally show ?thesis . qed lemma largest_eq: "largest TYPE(('e, 'f)float) = (2 ^ (LENGTH('f) + 1) - 1) * 2 powr real_of_int (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" proof - have "2 ^ LENGTH('e) - 1 - 1 = (2::nat) ^ LENGTH('e) - 2" by arith then have "largest TYPE(('e, 'f)float) = (2 ^ (LENGTH('f) + 1) - 1) * 2 powr (real (2 ^ LENGTH('e) - 2) + 1 - real (2 ^ (LENGTH('e) - 1)) - LENGTH('f))" unfolding largest_def emax_eq bias_def by (auto simp: largest_def powr_realpow[symmetric] powr_minus divide_simps algebra_simps powr_diff powr_add of_nat_diff) also have "2 ^ LENGTH('e) \ (2::nat)" by (simp add: self_le_power) then have "(real (2 ^ LENGTH('e) - 2) + 1 - real (2 ^ (LENGTH('e) - 1)) - LENGTH('f)) = (real (2 ^ LENGTH('e)) - 2 ^ (LENGTH('e) - 1) - LENGTH('f)) - 1" by (auto simp add: of_nat_diff) also have "real (2 ^ LENGTH('e)) = 2 ^ LENGTH('e)" by auto also have "(2 ^ LENGTH('e) - 2 ^ (LENGTH('e) - 1) - real LENGTH('f) - 1) = real_of_int ((2 ^ (LENGTH('e) - 1) - int (LENGTH('f)) - 1))" by (simp, subst power_Suc[symmetric], simp) finally show ?thesis by simp qed lemma bitlen_denormal_mantissa: "bitlen (abs (denormal_mantissa x)) \ LENGTH('f)" for x::"('e, 'f)float" proof - have "fraction x < 2 ^ LENGTH('f)" using float_frac_le[of x] by (metis One_nat_def Suc_pred le_imp_less_Suc pos2 zero_less_power) moreover have "- int (fraction x) \ 2 ^ LENGTH('f)" using negative_zle_0 order_trans zero_le_numeral zero_le_power by blast ultimately show ?thesis by (cases x rule: sign_cases) (auto simp: bitlen_le_iff_power denormal_mantissa_def intro!: antisym) qed lemma float_le_topfloat: fixes a::"('e, 'f)float" assumes "is_finite a" "LENGTH('e) > 1" shows "a \ topfloat" using assms(1) proof (induction a rule: float_induct) case (normal x) note normal(2) also have "real_of_int (normal_mantissa x) * 2 powr real_of_int (normal_exponent x) \ (2 powr (LENGTH('f) + 1) - 1) * 2 powr (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" using normal_exponent_bounds_real(2)[OF \is_normal x\] by (auto intro!: mult_mono real_of_int_le_2_powr_bitlenI simp: bitlen_normal_mantissa powr_realpow[symmetric] ge_one_powr_ge_zero) also have "\ = largest TYPE(('e, 'f) IEEE.float)" unfolding largest_eq by (auto simp: powr_realpow powr_add) also have "\ = valof (topfloat::('e, 'f) float)" using assms by (simp add: valof_topfloat) finally show ?case by (intro float_leI normal finite_topfloat) next case (denormal x) note denormal(2) also have "3 \ 2 powr (1 + real (LENGTH('e) - Suc 0))" proof - have "3 \ 2 powr (2::real)" by simp also have "\ \ 2 powr (1 + real (LENGTH('e) - Suc 0))" using assms by (subst powr_le_cancel_iff) auto finally show ?thesis . qed then have "real_of_int (denormal_mantissa x) * 2 powr real_of_int (denormal_exponent TYPE(('e, 'f)float)) \ (2 powr (LENGTH('f) + 1) - 1) * 2 powr (2 ^ (LENGTH('e) - 1) - int LENGTH('f) - 1)" using bitlen_denormal_mantissa[of x] by (auto intro!: mult_mono real_of_int_le_2_powr_bitlenI simp: bitlen_normal_mantissa powr_realpow[symmetric] ge_one_powr_ge_zero denormal_exponent_def bias_def powr_mult_base of_nat_diff) also have "\ \ largest TYPE(('e, 'f) IEEE.float)" unfolding largest_eq by (rule mult_mono) (auto simp: powr_realpow powr_add power_Suc[symmetric] simp del: power_Suc) also have "\ = valof (topfloat::('e, 'f) float)" using assms by (simp add: valof_topfloat) finally show ?case by (intro float_leI denormal finite_topfloat) qed auto lemma float_val_le_largest: "valof a \ largest TYPE(('e, 'f)float)" if "is_finite a" "LENGTH('e) > 1" for a::"('e, 'f)float" by (metis that finite_topfloat float_le float_le_topfloat valof_topfloat) lemma float_val_lt_threshold: "valof a < threshold TYPE(('e, 'f)float)" if "is_finite a" "LENGTH('e) > 1" for a::"('e, 'f)float" proof - have "valof a \ largest TYPE(('e, 'f)float)" by (rule float_val_le_largest [OF that]) also have "\ < threshold TYPE(('e, 'f)float)" by (auto simp: largest_def threshold_def divide_simps) finally show ?thesis . qed subsection \Algebraic properties about basic arithmetic\ text \Commutativity of addition.\ lemma assumes "is_finite a" "is_finite b" shows float_plus_comm_eq: "a + b = b + a" and float_plus_comm: "is_finite (a + b) \ (a + b) \ (b + a)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then show "a + b = b + a" by (simp add: float_defs fadd_def plus_float_def add.commute) then show "is_finite (a + b) \ (a + b) \ (b + a)" by (metis float_eq) qed text \The floating-point number a falls into the same category as the negation of \a\.\ lemma is_zero_uminus[simp]: "is_zero (- a) \ is_zero a" by (simp add: is_zero_def) lemma is_infinity_uminus [simp]: "is_infinity (- a) = is_infinity a" by (simp add: is_infinity_def) lemma is_finite_uminus[simp]: "is_finite (- a) \ is_finite a" by (simp add: is_finite_def) lemma is_nan_uminus[simp]: "is_nan (- a) \ is_nan a" by (simp add: is_nan_def) text \The sign of a and the sign of a's negation is different.\ lemma float_neg_sign: "(sign a) \ (sign (- a))" by (cases a rule: sign_cases) (auto simp: sign_minus_float) lemma float_neg_sign1: "sign a = sign (- b) \ sign a \ sign b" by (metis float_neg_sign sign_cases) lemma valof_uminus: assumes "is_finite a" shows "valof (- a) = - valof a" (is "?L = ?R") by (cases a rule: sign_cases) (auto simp: valof_eq sign_minus_float) text \Showing \a + (- b) = a - b\.\ lemma float_neg_add: "is_finite a \ is_finite b \ is_finite (a - b) \ valof a + valof (- b) = valof a - valof b" by (simp add: valof_uminus) lemma float_plus_minus: assumes "is_finite a" "is_finite b" "is_finite (a - b)" shows "(a + - b) \ (a - b)" proof - have nab: "\ is_nan a" "\ is_nan (- b)" "\ is_infinity a" "\ is_infinity (- b)" using assms by (auto simp: finite_nan finite_infinity) have "a - b = (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then (sign a) else 0) (round To_nearest (valof a - valof b)))" using nab by (auto simp: minus_float_def fsub_def) also have "\ = (zerosign (if is_zero a \ is_zero (- b) \ sign a = sign (- b) then sign a else 0) (round To_nearest (valof a + valof (- b))))" using assms by (simp add: float_neg_sign1 float_neg_add) also have "\ = a + - b" using nab by (auto simp: float_defs fadd_def plus_float_def) finally show ?thesis using assms by (metis float_eq) qed lemma finite_bottomfloat: "is_finite bottomfloat" by (simp add: finite_topfloat) lemma bottomfloat_eq_m_largest: "valof (bottomfloat::('e, 'f)float) = - largest TYPE(('e, 'f)float)" if "LENGTH('e) > 1" using that by (auto simp: valof_topfloat valof_uminus finite_topfloat) lemma float_val_ge_bottomfloat: "valof a \ valof (bottomfloat::('e, 'f)float)" if "LENGTH('e) > 1" "is_finite a" for a::"('e,'f)float" proof - have "- a \ topfloat" using that by (auto intro: float_le_topfloat) then show ?thesis using that by (auto simp: valof_uminus finite_topfloat) qed lemma float_ge_bottomfloat: "is_finite a \ a \ bottomfloat" if "LENGTH('e) > 1" "is_finite a" for a::"('e,'f)float" by (metis finite_bottomfloat float_le float_val_ge_bottomfloat that) lemma float_val_ge_largest: fixes a::"('e,'f)float" assumes "LENGTH('e) > 1" "is_finite a" shows "valof a \ - largest TYPE(('e,'f)float)" proof - have "- largest TYPE(('e,'f)float) = valof (bottomfloat::('e,'f)float)" using assms by (simp add: bottomfloat_eq_m_largest) also have "\ \ valof a" using assms by (simp add: float_val_ge_bottomfloat) finally show ?thesis . qed lemma float_val_gt_threshold: fixes a::"('e,'f)float" assumes "LENGTH('e) > 1" "is_finite a" shows "valof a > - threshold TYPE(('e,'f)float)" proof - have largest: "valof a \ -largest TYPE(('e,'f)float)" using assms by (metis float_val_ge_largest) then have "-largest TYPE(('e,'f)float) > - threshold TYPE(('e,'f)float)" by (auto simp: bias_def threshold_def largest_def divide_simps) then show ?thesis by (metis largest less_le_trans) qed text \Showing \abs (- a) = abs a\.\ lemma float_abs [simp]: "\ is_nan a \ abs (- a) = abs a" by (metis IEEE.abs_float_def float_neg_sign1 minus_minus_float zero_simps(1)) lemma neg_zerosign: "- (zerosign s a) = zerosign (1 - s) (- a)" by (auto simp: zerosign_def) subsection \Properties about Rounding Errors\ definition error :: "('e, 'f)float itself \ real \ real" where "error _ x = valof (round To_nearest x::('e, 'f)float) - x" lemma bound_at_worst_lemma: fixes a::"('e, 'f)float" assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" assumes finite: "is_finite a" shows "\valof (round To_nearest x::('e, 'f)float) - x\ \ \valof a - x\" proof - have *: "(round To_nearest x::('e,'f)float) = closest valof (\a. even (fraction a)) {a. is_finite a} x" using threshold finite by auto have "is_closest (valof) {a. is_finite a} x (round To_nearest x::('e,'f)float)" using is_finite_nonempty unfolding * by (intro closest_is_closest) auto then show ?thesis using finite is_closest_def by (metis mem_Collect_eq) qed lemma error_at_worst_lemma: fixes a::"('e, 'f)float" assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" and "is_finite a" shows "\error TYPE(('e, 'f)float) x\ \ \valof a - x\ " unfolding error_def by (rule bound_at_worst_lemma; fact) lemma error_is_zero [simp]: fixes a::"('e, 'f)float" assumes "is_finite a" "1 < LENGTH('e)" shows "error TYPE(('e, 'f)float) (valof a) = 0" proof - have "\valof a\ < threshold TYPE(('e, 'f)float)" by (metis abs_less_iff minus_less_iff float_val_gt_threshold float_val_lt_threshold assms) then show ?thesis by (metis abs_le_zero_iff abs_zero diff_self error_at_worst_lemma assms(1)) qed lemma is_finite_zerosign[simp]: "is_finite (zerosign s a) \ is_finite a" by (auto simp: zerosign_def is_finite_def) lemma is_finite_closest: "is_finite (closest (v::_\real) p {a. is_finite a} x)" using closest_is_closest[OF is_finite_nonempty, of v x p] by (auto simp: is_closest_def) lemma defloat_float_zerosign_round_finite: assumes threshold: "\x\ < threshold TYPE(('e, 'f)float)" shows "is_finite (zerosign s (round To_nearest x::('e, 'f)float))" proof - have "(round To_nearest x::('e, 'f)float) = (closest valof (\a. even (fraction a)) {a. is_finite a} x)" using threshold by (metis (full_types) abs_less_iff leD le_minus_iff round.simps(1)) then have "is_finite (round To_nearest x::('e, 'f)float)" by (metis is_finite_closest) then show ?thesis using is_finite_zerosign by auto qed lemma valof_zero[simp]: "valof 0 = 0" "valof minus_zero = 0" by (auto simp add: zerosign_def valof_eq zero_simps) lemma signzero_zero: "is_zero a \ valof (zerosign s a) = 0" by (auto simp add: zerosign_def) lemma val_zero: "is_zero a \ valof a = 0" by (cases a rule: is_zero_cases) auto lemma float_add: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a + valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_add: "is_finite (a + b)" and error_float_add: "valof (a + b) = valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms float_distinct_finite by auto then have ab: "(a + b) = (zerosign (if is_zero a \ is_zero b \ sign a = sign b then (sign a) else 0) (round To_nearest (valof a + valof b)))" using assms by (auto simp add: float_defs fadd_def plus_float_def) then show "is_finite ((a + b))" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a + b) = valof (zerosign (if is_zero a \ is_zero b \ sign a = sign b then (sign a) else 0) (round To_nearest (valof a + valof b)::('e, 'f)float))" by (auto simp: ab is_infinity_def is_nan_def valof_def) show "valof (a + b) = valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b)" proof (cases "is_zero (round To_nearest (valof a + valof b)::('e, 'f)float)") case True have "valof a + valof b + error TYPE(('e, 'f)float) (valof a + valof b) = valof (round To_nearest (valof a + valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_sub: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a - valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_sub: "is_finite (a - b)" and error_float_sub: "valof (a - b) = valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b)" proof - have "\ is_nan a" and "\ is_nan b" and "\ is_infinity a" and "\ is_infinity b" using assms by (auto simp: finite_nan finite_infinity) then have ab: "a - b = (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then sign a else 0) (round To_nearest (valof a - valof b)))" using assms by (auto simp add: float_defs fsub_def minus_float_def) then show "is_finite (a - b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a - b) = valof (zerosign (if is_zero a \ is_zero b \ sign a \ sign b then sign a else 0) (round To_nearest (valof a - valof b)::('e, 'f)float))" by (auto simp: ab is_infinity_def is_nan_def valof_def) show "valof (a - b) = valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b)" proof (cases "is_zero (round To_nearest (valof a - valof b)::('e, 'f)float)") case True have "valof a - valof b + error TYPE(('e, 'f)float) (valof a - valof b) = valof (round To_nearest (valof a - valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_mul: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and threshold: "\valof a * valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_mul: "is_finite (a * b)" and error_float_mul: "valof (a * b) = valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b)" proof - have non: "\ is_nan a" "\ is_nan b" "\ is_infinity a" "\ is_infinity b" using assms float_distinct_finite by auto then have ab: "a * b = (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a * valof b)::('e, 'f)float))" using assms by (auto simp: float_defs fmul_def times_float_def) then show "is_finite (a * b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a * b) = valof (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a * valof b)::('e, 'f)float))" by (auto simp: ab float_defs of_bool_def) show "valof (a * b) = valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b)" proof (cases "is_zero (round To_nearest (valof a * valof b)::('e, 'f)float)") case True have "valof a * valof b + error TYPE(('e, 'f)float) (valof a * valof b) = valof (round To_nearest (valof a * valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma float_div: fixes a b::"('e, 'f)float" assumes "is_finite a" and "is_finite b" and not_zero: "\ is_zero b" and threshold: "\valof a / valof b\ < threshold TYPE(('e, 'f)float)" shows finite_float_div: "is_finite (a / b)" and error_float_div: "valof (a / b) = valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b)" proof - have ab: "a / b = (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a / valof b)))" using assms by (simp add: divide_float_def fdiv_def finite_infinity finite_nan not_zero float_defs [symmetric]) then show "is_finite (a / b)" by (metis threshold defloat_float_zerosign_round_finite) have val_ab: "valof (a / b) = valof (zerosign (of_bool (sign a \ sign b)) (round To_nearest (valof a / valof b))::('e, 'f)float)" by (auto simp: ab float_defs of_bool_def) show "valof (a / b) = valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b)" proof (cases "is_zero (round To_nearest (valof a / valof b)::('e, 'f)float)") case True have "valof a / valof b + error TYPE(('e, 'f)float) (valof a / valof b) = valof (round To_nearest (valof a / valof b)::('e, 'f)float)" unfolding error_def by simp then show ?thesis by (metis True signzero_zero val_zero val_ab) next case False then show ?thesis by (metis ab add.commute eq_diff_eq' error_def zerosign_def) qed qed lemma valof_one[simp]: "valof (1::('e, 'f)float) = (if LENGTH('e) \ 1 then 0 else 1)" by transfer (auto simp: bias_def unat_sub word_1_le_power p2_eq_1) end diff --git a/thys/IP_Addresses/IP_Address.thy b/thys/IP_Addresses/IP_Address.thy --- a/thys/IP_Addresses/IP_Address.thy +++ b/thys/IP_Addresses/IP_Address.thy @@ -1,415 +1,415 @@ (* Title: IP_Address.thy Authors: Cornelius Diekmann *) theory IP_Address imports Word_Lib.Word_Lemmas Word_Lib.Word_Syntax Word_Lib.Reversed_Bit_Lists Hs_Compat WordInterval begin section \Modelling IP Adresses\ text\An IP address is basically an unsigned integer. We model IP addresses of arbitrary lengths. We will write @{typ "'i::len word"} for IP addresses of length @{term "LENGTH('i::len)"}. We use the convention to write @{typ 'i} whenever we mean IP addresses instead of generic words. When we will later have theorems with several polymorphic types in it (e.g. arbitrarily extensible packets), this notation makes it easier to spot that type @{typ 'i} is for IP addresses. The files @{file \IPv4.thy\} @{file \IPv6.thy\} concrete this for IPv4 and IPv6.\ text\The maximum IP address\ definition max_ip_addr :: "'i::len word" where "max_ip_addr \ of_nat ((2^(len_of(TYPE('i)))) - 1)" - lemma max_ip_addr_max_word: "max_ip_addr = max_word" + lemma max_ip_addr_max_word: "max_ip_addr = - 1" by (simp only: max_ip_addr_def of_nat_mask_eq flip: mask_eq_exp_minus_1) simp lemma max_ip_addr_max: "\a. a \ max_ip_addr" by(simp add: max_ip_addr_max_word) lemma range_0_max_UNIV: "UNIV = {0 .. max_ip_addr}" (*not in the simp set, for a reason*) by(simp add: max_ip_addr_max_word) fastforce lemma "size (x::'i::len word) = len_of(TYPE('i))" by(simp add:word_size) subsection\Sets of IP Addresses\ (*Warning, not executable!*) text\Specifying sets with network masks: 192.168.0.0 255.255.255.0\ definition ipset_from_netmask::"'i::len word \ 'i::len word \ 'i::len word set" where "ipset_from_netmask addr netmask \ let network_prefix = (addr AND netmask) in {network_prefix .. network_prefix OR (NOT netmask)}" text\Example (pseudo syntax): @{const ipset_from_netmask} \192.168.1.129 255.255.255.0\ = \{192.168.1.0 .. 192.168.1.255}\\ text\A network mask of all ones (i.e. @{term "(- 1)::'i::len word"}).\ lemma ipset_from_netmask_minusone: "ipset_from_netmask ip (- 1) = {ip}" by (simp add: ipset_from_netmask_def) lemma ipset_from_netmask_maxword: - "ipset_from_netmask ip max_word = {ip}" by (simp add: ipset_from_netmask_def) + "ipset_from_netmask ip (- 1) = {ip}" by (simp add: ipset_from_netmask_def) lemma ipset_from_netmask_zero: "ipset_from_netmask ip 0 = UNIV" by (auto simp add: ipset_from_netmask_def) text\Specifying sets in Classless Inter-domain Routing (CIDR) notation: 192.168.0.0/24\ definition ipset_from_cidr ::"'i::len word \ nat \ 'i::len word set" where "ipset_from_cidr addr pflength \ ipset_from_netmask addr ((mask pflength) << (len_of(TYPE('i)) - pflength))" text\Example (pseudo syntax): @{const ipset_from_cidr} \192.168.1.129 24\ = \{192.168.1.0 .. 192.168.1.255}\\ (*does this simplify stuff?*) lemma "(case ipcidr of (base, len) \ ipset_from_cidr base len) = uncurry ipset_from_cidr ipcidr" by(simp add: uncurry_case_stmt) lemma ipset_from_cidr_0: "ipset_from_cidr ip 0 = UNIV" by(auto simp add: ipset_from_cidr_def ipset_from_netmask_def Let_def) text\A prefix length of word size gives back the singleton set with the IP address. Example: \192.168.1.2/32 = {192.168.1.2}\\ lemma ipset_from_cidr_wordlength: fixes ip :: "'i::len word" shows "ipset_from_cidr ip (LENGTH('i)) = {ip}" by (simp add: ipset_from_cidr_def ipset_from_netmask_def) text\Alternative definition: Considering words as bit lists:\ lemma ipset_from_cidr_bl: fixes addr :: "'i::len word" shows "ipset_from_cidr addr pflength \ ipset_from_netmask addr (of_bl ((replicate pflength True) @ (replicate ((len_of(TYPE('i))) - pflength)) False))" by(simp add: ipset_from_cidr_def mask_bl shiftl_of_bl) lemma ipset_from_cidr_alt: fixes pre :: "'i::len word" shows "ipset_from_cidr pre len = {pre AND (mask len << LENGTH('i) - len) .. pre OR mask (LENGTH('i) - len)}" apply(simp add: ipset_from_cidr_def ipset_from_netmask_def Let_def) apply(simp add: word_oa_dist) apply(simp add: NOT_mask_shifted_lenword) done lemma ipset_from_cidr_alt2: fixes base ::"'i::len word" shows "ipset_from_cidr base len = ipset_from_netmask base (NOT (mask (LENGTH('i) - len)))" apply(simp add: ipset_from_cidr_def) using NOT_mask_shifted_lenword by(metis word_not_not) text\In CIDR notation, we cannot express the empty set.\ lemma ipset_from_cidr_not_empty: "ipset_from_cidr base len \ {}" by(simp add: ipset_from_cidr_alt bitmagic_zeroLast_leq_or1Last) text\Though we can write 192.168.1.2/24, we say that 192.168.0.0/24 is well-formed.\ lemma ipset_from_cidr_base_wellforemd: fixes base:: "'i::len word" assumes "mask (LENGTH('i) - l) AND base = 0" shows "ipset_from_cidr base l = {base .. base OR mask (LENGTH('i) - l)}" proof - have maskshift_eq_not_mask_generic: "((mask l << LENGTH('i) - l) :: 'i::len word) = NOT (mask (LENGTH('i) - l))" using NOT_mask_shifted_lenword by (metis word_not_not) have *: "base AND NOT (mask (LENGTH('i) - l)) = base" unfolding mask_eq_0_eq_x[symmetric] using assms word_bw_comms(1)[of base] by simp hence **: "base AND NOT (mask (LENGTH('i) - l)) OR mask (LENGTH('i) - l) = base OR mask (LENGTH('i) - l)" by simp have "ipset_from_netmask base (NOT (mask (LENGTH('i) - l))) = {base .. base || mask (LENGTH('i) - l)}" by(simp add: ipset_from_netmask_def Let_def ** *) thus ?thesis by(simp add: ipset_from_cidr_def maskshift_eq_not_mask_generic) qed lemma ipset_from_cidr_large_pfxlen: fixes ip:: "'i::len word" assumes "n \ LENGTH('i)" shows "ipset_from_cidr ip n = {ip}" proof - have obviously: "mask (LENGTH('i) - n) = 0" by (simp add: assms) show ?thesis apply(subst ipset_from_cidr_base_wellforemd) subgoal using assms by simp by (simp add: obviously) qed lemma ipset_from_netmask_base_mask_consume: fixes base :: "'i::len word" shows "ipset_from_netmask (base AND NOT (mask (LENGTH('i) - m))) (NOT (mask (LENGTH('i) - m))) = ipset_from_netmask base (NOT (mask (LENGTH('i) - m)))" unfolding ipset_from_netmask_def by(simp) text\Another definition of CIDR notation: All IP address which are equal on the first @{term "len - n"} bits\ definition ip_cidr_set :: "'i::len word \ nat \ 'i word set" where "ip_cidr_set i r \ {j . i AND NOT (mask (LENGTH('i) - r)) = j AND NOT (mask (LENGTH('i) - r))}" text\The definitions are equal\ lemma ipset_from_cidr_eq_ip_cidr_set: fixes base::"'i::len word" shows "ipset_from_cidr base len = ip_cidr_set base len" proof - have maskshift_eq_not_mask_generic: "((mask len << LENGTH('a) - len) :: 'a::len word) = NOT (mask (LENGTH('a) - len))" using NOT_mask_shifted_lenword by (metis word_not_not) have 1: "mask (len - m) AND base AND NOT (mask (len - m)) = 0" for len m and base::"'i::len word" by(simp add: word_bw_lcs) have 2: "mask (LENGTH('i) - len) AND pfxm_p = 0 \ (a \ ipset_from_netmask pfxm_p (NOT (mask (LENGTH('i) - len)))) \ (pfxm_p = NOT (mask (LENGTH('i) - len)) AND a)" for a::"'i::len word" and pfxm_p apply(subst ipset_from_cidr_alt2[symmetric]) apply(subst zero_base_lsb_imp_set_eq_as_bit_operation) apply(simp; fail) apply(subst ipset_from_cidr_base_wellforemd) apply(simp; fail) apply(simp) done from 2[OF 1, of _ base] have "(x \ ipset_from_netmask base (~~ (mask (LENGTH('i) - len)))) \ (base && ~~ (mask (LENGTH('i) - len)) = x && ~~ (mask (LENGTH('i) - len)))" for x apply(simp add: ipset_from_netmask_base_mask_consume) unfolding word_bw_comms(1)[of _ " ~~ (mask (LENGTH('i) - len))"] by simp then show ?thesis unfolding ip_cidr_set_def ipset_from_cidr_def by(auto simp add: maskshift_eq_not_mask_generic) qed lemma ip_cidr_set_change_base: "j \ ip_cidr_set i r \ ip_cidr_set j r = ip_cidr_set i r" by (auto simp: ip_cidr_set_def) subsection\IP Addresses as WordIntervals\ text\The nice thing is: @{typ "'i wordinterval"}s are executable.\ definition iprange_single :: "'i::len word \ 'i wordinterval" where "iprange_single ip \ WordInterval ip ip" fun iprange_interval :: "('i::len word \ 'i::len word) \ 'i wordinterval" where "iprange_interval (ip_start, ip_end) = WordInterval ip_start ip_end" declare iprange_interval.simps[simp del] lemma iprange_interval_uncurry: "iprange_interval ipcidr = uncurry WordInterval ipcidr" by(cases ipcidr) (simp add: iprange_interval.simps) lemma "wordinterval_to_set (iprange_single ip) = {ip}" by(simp add: iprange_single_def) lemma "wordinterval_to_set (iprange_interval (ip1, ip2)) = {ip1 .. ip2}" by(simp add: iprange_interval.simps) text\Now we can use the set operations on @{typ "'i::len wordinterval"}s\ term wordinterval_to_set term wordinterval_element term wordinterval_union term wordinterval_empty term wordinterval_setminus term wordinterval_UNIV term wordinterval_invert term wordinterval_intersection term wordinterval_subset term wordinterval_eq subsection\IP Addresses in CIDR Notation\ text\We want to convert IP addresses in CIDR notation to intervals. We already have @{const ipset_from_cidr}, which gives back a non-executable set. We want to convert to something we can store in an @{typ "'i wordinterval"}.\ fun ipcidr_to_interval_start :: "('i::len word \ nat) \ 'i::len word" where "ipcidr_to_interval_start (pre, len) = ( let netmask = (mask len) << (LENGTH('i) - len); network_prefix = (pre AND netmask) in network_prefix)" fun ipcidr_to_interval_end :: "('i::len word \ nat) \ 'i::len word" where "ipcidr_to_interval_end (pre, len) = ( let netmask = (mask len) << (LENGTH('i) - len); network_prefix = (pre AND netmask) in network_prefix OR (NOT netmask))" definition ipcidr_to_interval :: "('i::len word \ nat) \ ('i word \ 'i word)" where "ipcidr_to_interval cidr \ (ipcidr_to_interval_start cidr, ipcidr_to_interval_end cidr)" lemma ipset_from_cidr_ipcidr_to_interval: "ipset_from_cidr base len = {ipcidr_to_interval_start (base,len) .. ipcidr_to_interval_end (base,len)}" by(simp add: Let_def ipcidr_to_interval_def ipset_from_cidr_def ipset_from_netmask_def) declare ipcidr_to_interval_start.simps[simp del] ipcidr_to_interval_end.simps[simp del] lemma ipcidr_to_interval: "ipcidr_to_interval (base, len) = (s,e) \ ipset_from_cidr base len = {s .. e}" by (simp add: ipcidr_to_interval_def ipset_from_cidr_ipcidr_to_interval) definition ipcidr_tuple_to_wordinterval :: "('i::len word \ nat) \ 'i wordinterval" where "ipcidr_tuple_to_wordinterval iprng \ iprange_interval (ipcidr_to_interval iprng)" lemma wordinterval_to_set_ipcidr_tuple_to_wordinterval: "wordinterval_to_set (ipcidr_tuple_to_wordinterval (b, m)) = ipset_from_cidr b m" unfolding ipcidr_tuple_to_wordinterval_def ipset_from_cidr_ipcidr_to_interval ipcidr_to_interval_def by(simp add: iprange_interval.simps) lemma wordinterval_to_set_ipcidr_tuple_to_wordinterval_uncurry: "wordinterval_to_set (ipcidr_tuple_to_wordinterval ipcidr) = uncurry ipset_from_cidr ipcidr" by(cases ipcidr, simp add: wordinterval_to_set_ipcidr_tuple_to_wordinterval) definition ipcidr_union_set :: "('i::len word \ nat) set \ ('i word) set" where "ipcidr_union_set ips \ \(base, len) \ ips. ipset_from_cidr base len" lemma ipcidr_union_set_uncurry: "ipcidr_union_set ips = (\ ipcidr \ ips. uncurry ipset_from_cidr ipcidr)" by(simp add: ipcidr_union_set_def uncurry_case_stmt) subsection\Clever Operations on IP Addresses in CIDR Notation\ text\Intersecting two intervals may result in a new interval. Example: \{1..10} \ {5..20} = {5..10}\ Intersecting two IP address ranges represented as CIDR ranges results either in the empty set or the smaller of the two ranges. It will never create a new range. \ context begin (*contributed by Lars Noschinski*) private lemma less_and_not_mask_eq: fixes i :: "('a :: len) word" assumes "r2 \ r1" "i && ~~ (mask r2) = x && ~~ (mask r2)" shows "i && ~~ (mask r1) = x && ~~ (mask r1)" proof - have "i AND NOT (mask r1) = (i && ~~ (mask r2)) && ~~ (mask r1)" (is "_ = ?w && _") using \r2 \ r1\ by (simp add: and_not_mask_twice max_def) also have "?w = x && ~~ (mask r2)" by fact also have "\ && ~~ (mask r1) = x && ~~ (mask r1)" using \r2 \ r1\ by (simp add: and_not_mask_twice max_def) finally show ?thesis . qed lemma ip_cidr_set_less: fixes i :: "'i::len word" shows "r1 \ r2 \ ip_cidr_set i r2 \ ip_cidr_set i r1" unfolding ip_cidr_set_def apply auto apply (rule less_and_not_mask_eq[where ?r2.0="LENGTH('i) - r2"]) apply auto done private lemma ip_cidr_set_intersect_subset_helper: fixes i1 r1 i2 r2 assumes disj: "ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 \ {}" and "r1 \ r2" shows "ip_cidr_set i2 r2 \ ip_cidr_set i1 r1" proof - from disj obtain j where "j \ ip_cidr_set i1 r1" "j \ ip_cidr_set i2 r2" by auto with \r1 \ r2\ have "j \ ip_cidr_set j r1" "j \ ip_cidr_set j r1" using ip_cidr_set_change_base ip_cidr_set_less by blast+ show "ip_cidr_set i2 r2 \ ip_cidr_set i1 r1" proof fix i assume "i \ ip_cidr_set i2 r2" with \j \ ip_cidr_set i2 r2\ have "i \ ip_cidr_set j r2" using ip_cidr_set_change_base by auto also have "ip_cidr_set j r2 \ ip_cidr_set j r1" using \r1 \ r2\ ip_cidr_set_less by blast also have "\ = ip_cidr_set i1 r1" using \j \ ip_cidr_set i1 r1\ ip_cidr_set_change_base by blast finally show "i \ ip_cidr_set i1 r1" . qed qed lemma ip_cidr_set_notsubset_empty_inter: "\ ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 \ \ ip_cidr_set i2 r2 \ ip_cidr_set i1 r1 \ ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 = {}" apply(cases "r1 \ r2") subgoal using ip_cidr_set_intersect_subset_helper by blast apply(cases "r2 \ r1") subgoal using ip_cidr_set_intersect_subset_helper by blast by(simp) end lemma ip_cidr_intersect: "\ ipset_from_cidr b2 m2 \ ipset_from_cidr b1 m1 \ \ ipset_from_cidr b1 m1 \ ipset_from_cidr b2 m2 \ ipset_from_cidr b1 m1 \ ipset_from_cidr b2 m2 = {}" apply(simp add: ipset_from_cidr_eq_ip_cidr_set) using ip_cidr_set_notsubset_empty_inter by blast text\Computing the intersection of two IP address ranges in CIDR notation\ fun ipcidr_conjunct :: "('i::len word \ nat) \ ('i word \ nat) \ ('i word \ nat) option" where "ipcidr_conjunct (base1, m1) (base2, m2) = ( if ipset_from_cidr base1 m1 \ ipset_from_cidr base2 m2 = {} then None else if ipset_from_cidr base1 m1 \ ipset_from_cidr base2 m2 then Some (base1, m1) else Some (base2, m2) )" text\Intersecting with an address with prefix length zero always yields a non-empty result.\ lemma ipcidr_conjunct_any: "ipcidr_conjunct a (x,0) \ None" "ipcidr_conjunct (y,0) b \ None" apply(cases a, simp add: ipset_from_cidr_0 ipset_from_cidr_not_empty) by(cases b, simp add: ipset_from_cidr_0 ipset_from_cidr_not_empty) lemma ipcidr_conjunct_correct: "(case ipcidr_conjunct (b1, m1) (b2, m2) of Some (bx, mx) \ ipset_from_cidr bx mx | None \ {}) = (ipset_from_cidr b1 m1) \ (ipset_from_cidr b2 m2)" apply(simp split: if_split_asm) using ip_cidr_intersect by fast declare ipcidr_conjunct.simps[simp del] subsection\Code Equations\ text\Executable definition using word intervals\ lemma ipcidr_conjunct_word[code_unfold]: "ipcidr_conjunct ips1 ips2 = ( if wordinterval_empty (wordinterval_intersection (ipcidr_tuple_to_wordinterval ips1) (ipcidr_tuple_to_wordinterval ips2)) then None else if wordinterval_subset (ipcidr_tuple_to_wordinterval ips1) (ipcidr_tuple_to_wordinterval ips2) then Some ips1 else Some ips2 )" apply(simp) apply(cases ips1, cases ips2, rename_tac b1 m1 b2 m2, simp) apply(auto simp add: wordinterval_to_set_ipcidr_tuple_to_wordinterval ipcidr_conjunct.simps split: if_split_asm) done (*with the code_unfold lemma before, this works!*) lemma "ipcidr_conjunct (0::32 word,0) (8,1) = Some (8, 1)" by eval export_code ipcidr_conjunct checking SML text\making element check executable\ lemma addr_in_ipset_from_netmask_code[code_unfold]: "addr \ (ipset_from_netmask base netmask) \ (base AND netmask) \ addr \ addr \ (base AND netmask) OR (NOT netmask)" by(simp add: ipset_from_netmask_def Let_def) lemma addr_in_ipset_from_cidr_code[code_unfold]: "(addr::'i::len word) \ (ipset_from_cidr pre len) \ (pre AND ((mask len) << (LENGTH('i) - len))) \ addr \ addr \ pre OR (mask (LENGTH('i) - len))" unfolding ipset_from_cidr_alt by simp end diff --git a/thys/IP_Addresses/IPv4.thy b/thys/IP_Addresses/IPv4.thy --- a/thys/IP_Addresses/IPv4.thy +++ b/thys/IP_Addresses/IPv4.thy @@ -1,261 +1,261 @@ (* Title: IPv4.thy Authors: Cornelius Diekmann, Julius Michaelis *) theory IPv4 imports IP_Address NumberWang_IPv4 (* include "HOL-Library.Code_Target_Nat" if you need to work with actual numbers.*) begin lemma take_bit_word_beyond_length_eq: \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by transfer simp section \IPv4 Adresses\ text\An IPv4 address is basically a 32 bit unsigned integer.\ type_synonym ipv4addr = "32 word" text\Conversion between natural numbers and IPv4 adresses\ definition nat_of_ipv4addr :: "ipv4addr \ nat" where "nat_of_ipv4addr a = unat a" definition ipv4addr_of_nat :: "nat \ ipv4addr" where "ipv4addr_of_nat n = of_nat n" text\The maximum IPv4 addres\ definition max_ipv4_addr :: "ipv4addr" where "max_ipv4_addr \ ipv4addr_of_nat ((2^32) - 1)" lemma max_ipv4_addr_number: "max_ipv4_addr = 4294967295" unfolding max_ipv4_addr_def ipv4addr_of_nat_def by(simp) lemma "max_ipv4_addr = 0b11111111111111111111111111111111" by(fact max_ipv4_addr_number) - lemma max_ipv4_addr_max_word: "max_ipv4_addr = max_word" + lemma max_ipv4_addr_max_word: "max_ipv4_addr = - 1" by(simp add: max_ipv4_addr_number) lemma max_ipv4_addr_max[simp]: "\a. a \ max_ipv4_addr" by(simp add: max_ipv4_addr_max_word) lemma UNIV_ipv4addrset: "UNIV = {0 .. max_ipv4_addr}" (*not in the simp set, for a reason*) by(simp add: max_ipv4_addr_max_word) fastforce text\identity functions\ lemma nat_of_ipv4addr_ipv4addr_of_nat_mod: "nat_of_ipv4addr (ipv4addr_of_nat n) = n mod 2^32" by (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def unat_of_nat take_bit_eq_mod) lemma nat_of_ipv4addr_ipv4addr_of_nat: "\ n \ nat_of_ipv4addr max_ipv4_addr \ \ nat_of_ipv4addr (ipv4addr_of_nat n) = n" by (simp add: nat_of_ipv4addr_ipv4addr_of_nat_mod max_ipv4_addr_def) lemma ipv4addr_of_nat_nat_of_ipv4addr: "ipv4addr_of_nat (nat_of_ipv4addr addr) = addr" by(simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) subsection\Representing IPv4 Adresses (Syntax)\ fun ipv4addr_of_dotdecimal :: "nat \ nat \ nat \ nat \ ipv4addr" where "ipv4addr_of_dotdecimal (a,b,c,d) = ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a )" fun dotdecimal_of_ipv4addr :: "ipv4addr \ nat \ nat \ nat \ nat" where "dotdecimal_of_ipv4addr a = (nat_of_ipv4addr ((a >> 24) AND 0xFF), nat_of_ipv4addr ((a >> 16) AND 0xFF), nat_of_ipv4addr ((a >> 8) AND 0xFF), nat_of_ipv4addr (a AND 0xff))" declare ipv4addr_of_dotdecimal.simps[simp del] declare dotdecimal_of_ipv4addr.simps[simp del] text\Examples:\ lemma "ipv4addr_of_dotdecimal (192, 168, 0, 1) = 3232235521" by(simp add: ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def) (*could be solved by eval, but needs "HOL-Library.Code_Target_Nat"*) lemma "dotdecimal_of_ipv4addr 3232235521 = (192, 168, 0, 1)" by(simp add: dotdecimal_of_ipv4addr.simps nat_of_ipv4addr_def) text\a different notation for @{term ipv4addr_of_dotdecimal}\ lemma ipv4addr_of_dotdecimal_bit: "ipv4addr_of_dotdecimal (a,b,c,d) = (ipv4addr_of_nat a << 24) + (ipv4addr_of_nat b << 16) + (ipv4addr_of_nat c << 8) + ipv4addr_of_nat d" proof - have a: "(ipv4addr_of_nat a) << 24 = ipv4addr_of_nat (a * 16777216)" by(simp add: ipv4addr_of_nat_def shiftl_t2n) have b: "(ipv4addr_of_nat b) << 16 = ipv4addr_of_nat (b * 65536)" by(simp add: ipv4addr_of_nat_def shiftl_t2n) have c: "(ipv4addr_of_nat c) << 8 = ipv4addr_of_nat (c * 256)" by(simp add: ipv4addr_of_nat_def shiftl_t2n) have ipv4addr_of_nat_suc: "\x. ipv4addr_of_nat (Suc x) = word_succ (ipv4addr_of_nat (x))" by(simp add: ipv4addr_of_nat_def, metis Abs_fnat_hom_Suc of_nat_Suc) { fix x y have "ipv4addr_of_nat x + ipv4addr_of_nat y = ipv4addr_of_nat (x+y)" apply(induction x arbitrary: y) apply(simp add: ipv4addr_of_nat_def; fail) by(simp add: ipv4addr_of_nat_suc word_succ_p1) } from this a b c show ?thesis apply(simp add: ipv4addr_of_dotdecimal.simps) apply(rule arg_cong[where f=ipv4addr_of_nat]) apply(thin_tac _)+ by presburger qed lemma size_ipv4addr: "size (x::ipv4addr) = 32" by(simp add:word_size) lemma dotdecimal_of_ipv4addr_ipv4addr_of_dotdecimal: "\ a < 256; b < 256; c < 256; d < 256 \ \ dotdecimal_of_ipv4addr (ipv4addr_of_dotdecimal (a,b,c,d)) = (a,b,c,d)" proof - assume "a < 256" and "b < 256" and "c < 256" and "d < 256" note assms= \a < 256\ \b < 256\ \c < 256\ \d < 256\ hence a: "nat_of_ipv4addr ((ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) >> 24) AND mask 8) = a" apply (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp add: drop_bit_take_bit nat_take_bit_eq flip: take_bit_eq_mask) apply (simp add: drop_bit_eq_div take_bit_eq_mod) done have ipv4addr_of_nat_AND_mask8: "(ipv4addr_of_nat a) AND mask 8 = (ipv4addr_of_nat (a mod 256))" for a apply (simp add: ipv4addr_of_nat_def) apply transfer apply (simp flip: take_bit_eq_mask) apply (simp add: take_bit_eq_mod of_nat_mod) done from assms have b: "nat_of_ipv4addr ((ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) >> 16) AND mask 8) = b" apply (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp add: drop_bit_take_bit flip: take_bit_eq_mask) using div65536 apply (simp add: drop_bit_eq_div take_bit_eq_mod) done from assms have c: "nat_of_ipv4addr ((ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) >> 8) AND mask 8) = c" apply (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp add: drop_bit_take_bit flip: take_bit_eq_mask) using div256 apply (simp add: drop_bit_eq_div take_bit_eq_mod) done from \d < 256\ have d: "nat_of_ipv4addr (ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) AND mask 8) = d" apply (simp add: ipv4addr_of_nat_AND_mask8 ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp flip: take_bit_eq_mask) apply (simp add: take_bit_eq_mod nat_mod_distrib nat_add_distrib nat_mult_distrib mod256) done from a b c d show ?thesis apply (simp add: ipv4addr_of_dotdecimal.simps dotdecimal_of_ipv4addr.simps) apply (simp add: mask_eq) done qed lemma ipv4addr_of_dotdecimal_dotdecimal_of_ipv4addr: "(ipv4addr_of_dotdecimal (dotdecimal_of_ipv4addr ip)) = ip" proof - have ip_and_mask8_bl_drop24: "(ip::ipv4addr) AND mask 8 = of_bl (drop 24 (to_bl ip))" by(simp add: of_drop_to_bl size_ipv4addr) have List_rev_drop_geqn: "length x \ n \ (take n (rev x)) = rev (drop (length x - n) x)" for x :: "'a list" and n by(simp add: List.rev_drop) have and_mask_bl_take: "length x \ n \ ((of_bl x) AND mask n) = (of_bl (rev (take n (rev (x)))))" for x n by(simp add: List_rev_drop_geqn of_bl_drop) have ipv4addr_and_255: "x AND 255 = x AND mask 8" for x :: ipv4addr by(simp add: mask_eq) have bit_equality: "((ip >> 24) AND 0xFF << 24) + ((ip >> 16) AND 0xFF << 16) + ((ip >> 8) AND 0xFF << 8) + (ip AND 0xFF) = of_bl (take 8 (to_bl ip) @ take 8 (drop 8 (to_bl ip)) @ take 8 (drop 16 (to_bl ip)) @ drop 24 (to_bl ip))" apply(simp add: ipv4addr_and_255) apply(simp add: shiftr_slice) apply(simp add: slice_take' size_ipv4addr) apply(simp add: and_mask_bl_take) apply(simp add: List_rev_drop_geqn) apply(simp add: drop_take) apply(simp add: shiftl_of_bl) apply(simp add: of_bl_append) apply(simp add: ip_and_mask8_bl_drop24) done have blip_split: "\ blip. length blip = 32 \ blip = (take 8 blip) @ (take 8 (drop 8 blip)) @ (take 8 (drop 16 blip)) @ (take 8 (drop 24 blip))" by(rename_tac blip,case_tac blip,simp_all)+ (*I'm so sorry for this ...*) have "ipv4addr_of_dotdecimal (dotdecimal_of_ipv4addr ip) = of_bl (to_bl ip)" apply(subst blip_split) apply(simp; fail) apply(simp add: ipv4addr_of_dotdecimal_bit dotdecimal_of_ipv4addr.simps) apply(simp add: ipv4addr_of_nat_nat_of_ipv4addr) apply(simp add: bit_equality) done thus ?thesis using word_bl.Rep_inverse[symmetric] by simp qed lemma ipv4addr_of_dotdecimal_eqE: "\ ipv4addr_of_dotdecimal (a,b,c,d) = ipv4addr_of_dotdecimal (e,f,g,h); a < 256; b < 256; c < 256; d < 256; e < 256; f < 256; g < 256; h < 256 \ \ a = e \ b = f \ c = g \ d = h" by (metis Pair_inject dotdecimal_of_ipv4addr_ipv4addr_of_dotdecimal) subsection\IP Ranges: Examples\ lemma "(UNIV :: ipv4addr set) = {0 .. max_ipv4_addr}" by(simp add: UNIV_ipv4addrset) lemma "(42::ipv4addr) \ UNIV" by(simp) (*Warning, not executable!*) lemma "ipset_from_netmask (ipv4addr_of_dotdecimal (192,168,0,42)) (ipv4addr_of_dotdecimal (255,255,0,0)) = {ipv4addr_of_dotdecimal (192,168,0,0) .. ipv4addr_of_dotdecimal (192,168,255,255)}" by(simp add: ipset_from_netmask_def ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def) lemma "ipset_from_netmask (ipv4addr_of_dotdecimal (192,168,0,42)) (ipv4addr_of_dotdecimal (0,0,0,0)) = UNIV" by(simp add: UNIV_ipv4addrset ipset_from_netmask_def ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def max_ipv4_addr_max_word) text\192.168.0.0/24\ lemma fixes addr :: ipv4addr shows "ipset_from_cidr addr pflength = ipset_from_netmask addr ((mask pflength) << (32 - pflength))" by(simp add: ipset_from_cidr_def) lemma "ipset_from_cidr (ipv4addr_of_dotdecimal (192,168,0,42)) 16 = {ipv4addr_of_dotdecimal (192,168,0,0) .. ipv4addr_of_dotdecimal (192,168,255,255)}" by(simp add: ipset_from_cidr_alt mask_eq ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def) lemma "ip \ (ipset_from_cidr (ipv4addr_of_dotdecimal (0, 0, 0, 0)) 0)" by(simp add: ipset_from_cidr_0) lemma ipv4set_from_cidr_32: fixes addr :: ipv4addr shows "ipset_from_cidr addr 32 = {addr}" by (simp add: ipset_from_cidr_alt take_bit_word_beyond_length_eq flip: take_bit_eq_mask) lemma fixes pre :: ipv4addr shows "ipset_from_cidr pre len = {(pre AND ((mask len) << (32 - len))) .. pre OR (mask (32 - len))}" by (simp add: ipset_from_cidr_alt ipset_from_cidr_def) text\making element check executable\ lemma addr_in_ipv4set_from_netmask_code[code_unfold]: fixes addr :: ipv4addr shows "addr \ (ipset_from_netmask base netmask) \ (base AND netmask) \ addr \ addr \ (base AND netmask) OR (NOT netmask)" by (simp add: addr_in_ipset_from_netmask_code) lemma addr_in_ipv4set_from_cidr_code[code_unfold]: fixes addr :: ipv4addr shows "addr \ (ipset_from_cidr pre len) \ (pre AND ((mask len) << (32 - len))) \ addr \ addr \ pre OR (mask (32 - len))" by(simp add: addr_in_ipset_from_cidr_code) (*small numbers because we didn't load Code_Target_Nat. Should work by eval*) lemma "ipv4addr_of_dotdecimal (192,168,42,8) \ (ipset_from_cidr (ipv4addr_of_dotdecimal (192,168,0,0)) 16)" by (simp add: ipv4addr_of_nat_def ipset_from_cidr_def ipv4addr_of_dotdecimal.simps ipset_from_netmask_def mask_eq_exp_minus_1 word_le_def take_bit_minus_one_eq_mask) definition ipv4range_UNIV :: "32 wordinterval" where "ipv4range_UNIV \ wordinterval_UNIV" lemma ipv4range_UNIV_set_eq: "wordinterval_to_set ipv4range_UNIV = UNIV" by(simp only: ipv4range_UNIV_def wordinterval_UNIV_set_eq) thm iffD1[OF wordinterval_eq_set_eq] (*TODO: probably the following is a good idea?*) (* declare iffD1[OF wordinterval_eq_set_eq, cong] *) text\This \LENGTH('a)\ is 32 for IPv4 addresses.\ lemma ipv4cidr_to_interval_simps[code_unfold]: "ipcidr_to_interval ((pre::ipv4addr), len) = ( let netmask = (mask len) << (32 - len); network_prefix = (pre AND netmask) in (network_prefix, network_prefix OR (NOT netmask)))" by(simp add: ipcidr_to_interval_def Let_def ipcidr_to_interval_start.simps ipcidr_to_interval_end.simps) end diff --git a/thys/IP_Addresses/IPv6.thy b/thys/IP_Addresses/IPv6.thy --- a/thys/IP_Addresses/IPv6.thy +++ b/thys/IP_Addresses/IPv6.thy @@ -1,946 +1,946 @@ (* Title: IPv6.thy Authors: Cornelius Diekmann *) theory IPv6 imports IP_Address NumberWang_IPv6 (* include "HOL-Library.Code_Target_Nat" if you need to work with actual numbers.*) 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 = max_word" + 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}" (*not in the simp set, for a reason*) 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} \ datatype ipv6addr_syntax = IPv6AddrPreferred "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" text_raw\ \begin{verbatim} 2. [...] In order to make writing addresses containing zero bits easier, a special syntax is available to compress the zeros. The use of "::" indicates one or more groups of 16 bits of zeros. The "::" can only appear once in an address. The "::" can also be used to compress leading or trailing zeros in an address. For example, the following addresses 2001:DB8:0:0:8:800:200C:417A a unicast address FF01:0:0:0:0:0:0:101 a multicast address 0:0:0:0:0:0:0:1 the loopback address 0:0:0:0:0:0:0:0 the unspecified address may be represented as 2001:DB8::8:800:200C:417A a unicast address FF01::101 a multicast address ::1 the loopback address :: the unspecified address \end{verbatim} \ (*datatype may take some minutes to load*) datatype ipv6addr_syntax_compressed = \ \using @{typ unit} for the omission @{text "::"}. Naming convention of the datatype: The first number is the position where the omission occurs. The second number is the length of the specified address pieces. I.e. `8 minus the second number' pieces are omitted.\ IPv6AddrCompressed1_0 unit | IPv6AddrCompressed1_1 unit "16 word" | IPv6AddrCompressed1_2 unit "16 word" "16 word" | IPv6AddrCompressed1_3 unit "16 word" "16 word" "16 word" | IPv6AddrCompressed1_4 unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_5 unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_6 unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_7 unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_1 "16 word" unit | IPv6AddrCompressed2_2 "16 word" unit "16 word" | IPv6AddrCompressed2_3 "16 word" unit "16 word" "16 word" | IPv6AddrCompressed2_4 "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed2_5 "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_6 "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_7 "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed3_2 "16 word" "16 word" unit | IPv6AddrCompressed3_3 "16 word" "16 word" unit "16 word" | IPv6AddrCompressed3_4 "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed3_5 "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed3_6 "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed3_7 "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed4_3 "16 word" "16 word" "16 word" unit | IPv6AddrCompressed4_4 "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed4_5 "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed4_6 "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed4_7 "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed5_4 "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed5_5 "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed5_6 "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed5_7 "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed6_5 "16 word" "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed6_6 "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed6_7 "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed7_6 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed7_7 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed8_7 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit (*RFC 5952: """ 4. A Recommendation for IPv6 Text Representation 4.2.2. Handling One 16-Bit 0 Field The symbol "::" MUST NOT be used to shorten just one 16-bit 0 field. For example, the representation 2001:db8:0:1:1:1:1:1 is correct, but 2001:db8::1:1:1:1:1 is not correct. """ So we could remove all IPv6AddrCompressed*_7 constructors. But these are `recommendations', we might still see these non-recommended definitions. "[...] all implementations must accept and be able to handle any legitimate RFC 4291 format." *) (*More convenient parser helper function for compressed IPv6 addresses: Input list (from parser): Some 16word \ address piece None \ omission '::' Basically, the parser must only do the following (python syntax): split the string which is an ipv6 address at ':' map empty string to None map everything else to Some (string_to_16word str) sanitize empty strings at the start and the end (see toString and parser theories) Example: "1:2:3".split(":") = ['1', '2', '3'] ":2:3:4".split(":") = ['', '2', '3', '4'] ":2::3".split(":") = ['', '2', '', '3'] "1:2:3:".split(":") = ['1', '2', '3', ''] *) definition parse_ipv6_address_compressed :: "((16 word) option) list \ ipv6addr_syntax_compressed option" where "parse_ipv6_address_compressed as = (case as of [None] \ Some (IPv6AddrCompressed1_0 ()) | [None, Some a] \ Some (IPv6AddrCompressed1_1 () a) | [None, Some a, Some b] \ Some (IPv6AddrCompressed1_2 () a b) | [None, Some a, Some b, Some c] \ Some (IPv6AddrCompressed1_3 () a b c) | [None, Some a, Some b, Some c, Some d] \ Some (IPv6AddrCompressed1_4 () a b c d) | [None, Some a, Some b, Some c, Some d, Some e] \ Some (IPv6AddrCompressed1_5 () a b c d e) | [None, Some a, Some b, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed1_6 () a b c d e f) | [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed1_7 () a b c d e f g) | [Some a, None] \ Some (IPv6AddrCompressed2_1 a ()) | [Some a, None, Some b] \ Some (IPv6AddrCompressed2_2 a () b) | [Some a, None, Some b, Some c] \ Some (IPv6AddrCompressed2_3 a () b c) | [Some a, None, Some b, Some c, Some d] \ Some (IPv6AddrCompressed2_4 a () b c d) | [Some a, None, Some b, Some c, Some d, Some e] \ Some (IPv6AddrCompressed2_5 a () b c d e) | [Some a, None, Some b, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed2_6 a () b c d e f) | [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed2_7 a () b c d e f g) | [Some a, Some b, None] \ Some (IPv6AddrCompressed3_2 a b ()) | [Some a, Some b, None, Some c] \ Some (IPv6AddrCompressed3_3 a b () c) | [Some a, Some b, None, Some c, Some d] \ Some (IPv6AddrCompressed3_4 a b () c d) | [Some a, Some b, None, Some c, Some d, Some e] \ Some (IPv6AddrCompressed3_5 a b () c d e) | [Some a, Some b, None, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed3_6 a b () c d e f) | [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed3_7 a b () c d e f g) | [Some a, Some b, Some c, None] \ Some (IPv6AddrCompressed4_3 a b c ()) | [Some a, Some b, Some c, None, Some d] \ Some (IPv6AddrCompressed4_4 a b c () d) | [Some a, Some b, Some c, None, Some d, Some e] \ Some (IPv6AddrCompressed4_5 a b c () d e) | [Some a, Some b, Some c, None, Some d, Some e, Some f] \ Some (IPv6AddrCompressed4_6 a b c () d e f) | [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed4_7 a b c () d e f g) | [Some a, Some b, Some c, Some d, None] \ Some (IPv6AddrCompressed5_4 a b c d ()) | [Some a, Some b, Some c, Some d, None, Some e] \ Some (IPv6AddrCompressed5_5 a b c d () e) | [Some a, Some b, Some c, Some d, None, Some e, Some f] \ Some (IPv6AddrCompressed5_6 a b c d () e f) | [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g] \ Some (IPv6AddrCompressed5_7 a b c d () e f g) | [Some a, Some b, Some c, Some d, Some e, None] \ Some (IPv6AddrCompressed6_5 a b c d e ()) | [Some a, Some b, Some c, Some d, Some e, None, Some f] \ Some (IPv6AddrCompressed6_6 a b c d e () f) | [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g] \ Some (IPv6AddrCompressed6_7 a b c d e () f g) | [Some a, Some b, Some c, Some d, Some e, Some f, None] \ Some (IPv6AddrCompressed7_6 a b c d e f ()) | [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g] \ Some (IPv6AddrCompressed7_7 a b c d e f () g) | [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None] \ Some (IPv6AddrCompressed8_7 a b c d e f g ()) | _ \ None \ \invalid ipv6 copressed address.\ )" fun ipv6addr_syntax_compressed_to_list :: "ipv6addr_syntax_compressed \ ((16 word) option) list" where "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_0 _) = [None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_1 () a) = [None, Some a]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_2 () a b) = [None, Some a, Some b]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_3 () a b c) = [None, Some a, Some b, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_4 () a b c d) = [None, Some a, Some b, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_5 () a b c d e) = [None, Some a, Some b, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_6 () a b c d e f) = [None, Some a, Some b, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_7 () a b c d e f g) = [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_1 a ()) = [Some a, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_2 a () b) = [Some a, None, Some b]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_3 a () b c) = [Some a, None, Some b, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_4 a () b c d) = [Some a, None, Some b, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_5 a () b c d e) = [Some a, None, Some b, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_6 a () b c d e f) = [Some a, None, Some b, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_7 a () b c d e f g) = [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_2 a b ()) = [Some a, Some b, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_3 a b () c) = [Some a, Some b, None, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_4 a b () c d) = [Some a, Some b, None, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_5 a b () c d e) = [Some a, Some b, None, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_6 a b () c d e f) = [Some a, Some b, None, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_7 a b () c d e f g) = [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_3 a b c ()) = [Some a, Some b, Some c, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_4 a b c () d) = [Some a, Some b, Some c, None, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_5 a b c () d e) = [Some a, Some b, Some c, None, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_6 a b c () d e f) = [Some a, Some b, Some c, None, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_7 a b c () d e f g) = [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_4 a b c d ()) = [Some a, Some b, Some c, Some d, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_5 a b c d () e) = [Some a, Some b, Some c, Some d, None, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_6 a b c d () e f) = [Some a, Some b, Some c, Some d, None, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_7 a b c d () e f g) = [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_5 a b c d e ()) = [Some a, Some b, Some c, Some d, Some e, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_6 a b c d e () f) = [Some a, Some b, Some c, Some d, Some e, None, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_7 a b c d e () f g) = [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed7_6 a b c d e f ()) = [Some a, Some b, Some c, Some d, Some e, Some f, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed7_7 a b c d e f () g) = [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed8_7 a b c d e f g ()) = [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None]" (*for all ipv6_syntax, there is a corresponding list representation*) lemma parse_ipv6_address_compressed_exists: obtains ss where "parse_ipv6_address_compressed ss = Some ipv6_syntax" proof define ss where "ss = ipv6addr_syntax_compressed_to_list ipv6_syntax" thus "parse_ipv6_address_compressed ss = Some ipv6_syntax" by (cases ipv6_syntax; simp add: parse_ipv6_address_compressed_def) qed lemma parse_ipv6_address_compressed_identity: "parse_ipv6_address_compressed (ipv6addr_syntax_compressed_to_list (ipv6_syntax)) = Some ipv6_syntax" by(cases ipv6_syntax; simp add: parse_ipv6_address_compressed_def) lemma parse_ipv6_address_compressed_someE: assumes "parse_ipv6_address_compressed as = Some ipv6" obtains "as = [None]" "ipv6 = (IPv6AddrCompressed1_0 ())" | a where "as = [None, Some a]" "ipv6 = (IPv6AddrCompressed1_1 () a)" | a b where "as = [None, Some a, Some b]" "ipv6 = (IPv6AddrCompressed1_2 () a b)" | a b c where "as = [None, Some a, Some b, Some c]" "ipv6 = (IPv6AddrCompressed1_3 () a b c)" | a b c d where "as = [None, Some a, Some b, Some c, Some d]" "ipv6 = (IPv6AddrCompressed1_4 () a b c d)" | a b c d e where "as = [None, Some a, Some b, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed1_5 () a b c d e)" | a b c d e f where "as = [None, Some a, Some b, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed1_6 () a b c d e f)" | a b c d e f g where "as = [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed1_7 () a b c d e f g)" | a where "as = [Some a, None]" "ipv6 = (IPv6AddrCompressed2_1 a ())" | a b where "as = [Some a, None, Some b]" "ipv6 = (IPv6AddrCompressed2_2 a () b)" | a b c where "as = [Some a, None, Some b, Some c]" "ipv6 = (IPv6AddrCompressed2_3 a () b c)" | a b c d where "as = [Some a, None, Some b, Some c, Some d]" "ipv6 = (IPv6AddrCompressed2_4 a () b c d)" | a b c d e where "as = [Some a, None, Some b, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed2_5 a () b c d e)" | a b c d e f where "as = [Some a, None, Some b, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed2_6 a () b c d e f)" | a b c d e f g where "as = [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed2_7 a () b c d e f g)" | a b where "as = [Some a, Some b, None]" "ipv6 = (IPv6AddrCompressed3_2 a b ())" | a b c where "as = [Some a, Some b, None, Some c]" "ipv6 = (IPv6AddrCompressed3_3 a b () c)" | a b c d where "as = [Some a, Some b, None, Some c, Some d]" "ipv6 = (IPv6AddrCompressed3_4 a b () c d)" | a b c d e where "as = [Some a, Some b, None, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed3_5 a b () c d e)" | a b c d e f where "as = [Some a, Some b, None, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed3_6 a b () c d e f)" | a b c d e f g where "as = [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed3_7 a b () c d e f g)" | a b c where "as = [Some a, Some b, Some c, None]" "ipv6 = (IPv6AddrCompressed4_3 a b c ())" | a b c d where "as = [Some a, Some b, Some c, None, Some d]" "ipv6 = (IPv6AddrCompressed4_4 a b c () d)" | a b c d e where "as = [Some a, Some b, Some c, None, Some d, Some e]" "ipv6 = (IPv6AddrCompressed4_5 a b c () d e)" | a b c d e f where "as = [Some a, Some b, Some c, None, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed4_6 a b c () d e f)" | a b c d e f g where "as = [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed4_7 a b c () d e f g)" | a b c d where "as = [Some a, Some b, Some c, Some d, None]" "ipv6 = (IPv6AddrCompressed5_4 a b c d ())" | a b c d e where "as = [Some a, Some b, Some c, Some d, None, Some e]" "ipv6 = (IPv6AddrCompressed5_5 a b c d () e)" | a b c d e f where "as = [Some a, Some b, Some c, Some d, None, Some e, Some f]" "ipv6 = (IPv6AddrCompressed5_6 a b c d () e f)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed5_7 a b c d () e f g)" | a b c d e where "as = [Some a, Some b, Some c, Some d, Some e, None]" "ipv6 = (IPv6AddrCompressed6_5 a b c d e ())" | a b c d e f where "as = [Some a, Some b, Some c, Some d, Some e, None, Some f]" "ipv6 = (IPv6AddrCompressed6_6 a b c d e () f)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g]" "ipv6 = (IPv6AddrCompressed6_7 a b c d e () f g)" | a b c d e f where "as = [Some a, Some b, Some c, Some d, Some e, Some f, None]" "ipv6 = (IPv6AddrCompressed7_6 a b c d e f ())" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g]" "ipv6 = (IPv6AddrCompressed7_7 a b c d e f () g)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None]" "ipv6 = (IPv6AddrCompressed8_7 a b c d e f g ())" using assms unfolding parse_ipv6_address_compressed_def by (auto split: list.split_asm option.split_asm) (* takes a minute *) lemma parse_ipv6_address_compressed_identity2: "ipv6addr_syntax_compressed_to_list ipv6_syntax = ls \ (parse_ipv6_address_compressed ls) = Some ipv6_syntax" (is "?lhs = ?rhs") proof assume ?rhs thus ?lhs by (auto elim: parse_ipv6_address_compressed_someE) next assume ?lhs thus ?rhs by (cases ipv6_syntax) (auto simp: parse_ipv6_address_compressed_def) qed text\Valid IPv6 compressed notation: \<^item> at most one omission \<^item> at most 7 pieces \ lemma RFC_4291_format: "parse_ipv6_address_compressed as \ None \ length (filter (\p. p = None) as) = 1 \ length (filter (\p. p \ None) as) \ 7" (is "?lhs = ?rhs") proof assume ?lhs then obtain addr where "parse_ipv6_address_compressed as = Some addr" by blast thus ?rhs by (elim parse_ipv6_address_compressed_someE; simp) next assume ?rhs thus ?lhs unfolding parse_ipv6_address_compressed_def by (auto split: option.split list.split if_split_asm) qed text_raw\ \begin{verbatim} 3. An alternative form that is sometimes more convenient when dealing with a mixed environment of IPv4 and IPv6 nodes is x:x:x:x:x:x:d.d.d.d, where the 'x's are the hexadecimal values of the six high-order 16-bit pieces of the address, and the 'd's are the decimal values of the four low-order 8-bit pieces of the address (standard IPv4 representation). Examples: 0:0:0:0:0:0:13.1.68.3 0:0:0:0:0:FFFF:129.144.52.38 or in compressed form: ::13.1.68.3 ::FFFF:129.144.52.38 \end{verbatim} This is currently not supported by our library! \ (*TODO*) (*TODO: oh boy, they can also be compressed*) subsection\Semantics\ fun ipv6preferred_to_int :: "ipv6addr_syntax \ ipv6addr" where "ipv6preferred_to_int (IPv6AddrPreferred a b c d e f g h) = (ucast a << (16 * 7)) OR (ucast b << (16 * 6)) OR (ucast c << (16 * 5)) OR (ucast d << (16 * 4)) OR (ucast e << (16 * 3)) OR (ucast f << (16 * 2)) OR (ucast g << (16 * 1)) OR (ucast h << (16 * 0))" lemma "ipv6preferred_to_int (IPv6AddrPreferred 0x2001 0xDB8 0x0 0x0 0x8 0x800 0x200C 0x417A) = 42540766411282592856906245548098208122" by eval lemma "ipv6preferred_to_int (IPv6AddrPreferred 0xFF01 0x0 0x0 0x0 0x0 0x0 0x0 0x101) = 338958331222012082418099330867817087233" by eval declare ipv6preferred_to_int.simps[simp del] definition int_to_ipv6preferred :: "ipv6addr \ ipv6addr_syntax" where "int_to_ipv6preferred i = IPv6AddrPreferred (ucast ((i AND 0xFFFF0000000000000000000000000000) >> 16*7)) (ucast ((i AND 0xFFFF000000000000000000000000) >> 16*6)) (ucast ((i AND 0xFFFF00000000000000000000) >> 16*5)) (ucast ((i AND 0xFFFF0000000000000000) >> 16*4)) (ucast ((i AND 0xFFFF000000000000) >> 16*3)) (ucast ((i AND 0xFFFF00000000) >> 16*2)) (ucast ((i AND 0xFFFF0000) >> 16*1)) (ucast ((i AND 0xFFFF)))" lemma "int_to_ipv6preferred 42540766411282592856906245548098208122 = IPv6AddrPreferred 0x2001 0xDB8 0x0 0x0 0x8 0x800 0x200C 0x417A" by eval lemma word128_masks_ipv6pieces: "(0xFFFF0000000000000000000000000000::ipv6addr) = (mask 16) << 112" "(0xFFFF000000000000000000000000::ipv6addr) = (mask 16) << 96" "(0xFFFF00000000000000000000::ipv6addr) = (mask 16) << 80" "(0xFFFF0000000000000000::ipv6addr) = (mask 16) << 64" "(0xFFFF000000000000::ipv6addr) = (mask 16) << 48" "(0xFFFF00000000::ipv6addr) = (mask 16) << 32" "(0xFFFF0000::ipv6addr) = (mask 16) << 16" "(0xFFFF::ipv6addr) = (mask 16)" by(simp add: mask_eq)+ text\Correctness: round trip property one\ lemma ipv6preferred_to_int_int_to_ipv6preferred: "ipv6preferred_to_int (int_to_ipv6preferred ip) = ip" proof - have and_mask_shift_helper: "w AND (mask m << n) >> n << n = w AND (mask m << n)" for m n::nat and w::ipv6addr by (metis is_aligned_shift is_aligned_shiftr_shiftl shiftr_and_eq_shiftl) have ucast_ipv6_piece_rule: "length (dropWhile Not (to_bl w)) \ 16 \ (ucast::16 word \ 128 word) ((ucast::128 word \ 16 word) w) = w" for w::ipv6addr by(rule ucast_short_ucast_long_ingoreLeadingZero) (simp_all) have ucast_ipv6_piece: "16 \ 128 - n \ (ucast::16 word \ 128 word) ((ucast::128 word \ 16 word) (w AND (mask 16 << n) >> n)) << n = w AND (mask 16 << n)" for w::ipv6addr and n::nat apply(subst ucast_ipv6_piece_rule) apply(rule length_drop_mask_inner) apply(simp; fail) apply(subst and_mask_shift_helper) apply simp done have ucast16_ucast128_masks_highest_bits: "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000000000000000000000000000 >> 112)) << 112) = (ip AND 0xFFFF0000000000000000000000000000)" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF000000000000000000000000 >> 96)) << 96) = ip AND 0xFFFF000000000000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF00000000000000000000 >> 80)) << 80) = ip AND 0xFFFF00000000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000000000000000 >> 64)) << 64) = ip AND 0xFFFF0000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF000000000000 >> 48)) << 48) = ip AND 0xFFFF000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF00000000 >> 32)) << 32) = ip AND 0xFFFF00000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000 >> 16)) << 16) = ip AND 0xFFFF0000" by((subst word128_masks_ipv6pieces)+, subst ucast_ipv6_piece, simp_all)+ have ucast16_ucast128_masks_highest_bits0: "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF))) = ip AND 0xFFFF" apply(subst word128_masks_ipv6pieces)+ apply(subst ucast_short_ucast_long_ingoreLeadingZero) apply simp_all by (simp add: length_drop_mask) have mask_len_word:"n = (LENGTH('a)) \ w AND mask n = w" for n and w::"'a::len word" by (simp add: mask_eq_iff) have ipv6addr_16word_pieces_compose_or: "ip && (mask 16 << 112) || ip && (mask 16 << 96) || ip && (mask 16 << 80) || ip && (mask 16 << 64) || ip && (mask 16 << 48) || ip && (mask 16 << 32) || ip && (mask 16 << 16) || ip && mask 16 = ip" apply(subst word_ao_dist2[symmetric])+ apply(simp add: mask_eq) apply(subst mask128) apply(rule mask_len_word) apply simp done show ?thesis apply(simp add: ipv6preferred_to_int.simps int_to_ipv6preferred_def) apply(simp add: ucast16_ucast128_masks_highest_bits ucast16_ucast128_masks_highest_bits0) apply(simp add: word128_masks_ipv6pieces) apply(rule ipv6addr_16word_pieces_compose_or) done qed text\Correctness: round trip property two\ lemma int_to_ipv6preferred_ipv6preferred_to_int: "int_to_ipv6preferred (ipv6preferred_to_int ip) = ip" proof - note ucast_shift_simps=helper_masked_ucast_generic helper_masked_ucast_reverse_generic helper_masked_ucast_generic[where n=0, simplified] helper_masked_ucast_equal_generic note ucast_simps=helper_masked_ucast_reverse_generic[where m=0, simplified] helper_masked_ucast_equal_generic[where n=0, simplified] show ?thesis apply(cases ip, rename_tac a b c d e f g h) apply(simp add: ipv6preferred_to_int.simps int_to_ipv6preferred_def) apply(simp add: word128_masks_ipv6pieces) apply(simp add: word_ao_dist ucast_shift_simps ucast_simps) done qed text\compressed to preferred format\ fun ipv6addr_c2p :: "ipv6addr_syntax_compressed \ ipv6addr_syntax" where "ipv6addr_c2p (IPv6AddrCompressed1_0 ()) = IPv6AddrPreferred 0 0 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed1_1 () h) = IPv6AddrPreferred 0 0 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed1_2 () g h) = IPv6AddrPreferred 0 0 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed1_3 () f g h) = IPv6AddrPreferred 0 0 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_4 () e f g h) = IPv6AddrPreferred 0 0 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_5 () d e f g h) = IPv6AddrPreferred 0 0 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_6 () c d e f g h) = IPv6AddrPreferred 0 0 c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_7 () b c d e f g h) = IPv6AddrPreferred 0 b c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_1 a ()) = IPv6AddrPreferred a 0 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed2_2 a () h) = IPv6AddrPreferred a 0 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed2_3 a () g h) = IPv6AddrPreferred a 0 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed2_4 a () f g h) = IPv6AddrPreferred a 0 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_5 a () e f g h) = IPv6AddrPreferred a 0 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_6 a () d e f g h) = IPv6AddrPreferred a 0 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_7 a () c d e f g h) = IPv6AddrPreferred a 0 c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_2 a b ()) = IPv6AddrPreferred a b 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed3_3 a b () h) = IPv6AddrPreferred a b 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed3_4 a b () g h) = IPv6AddrPreferred a b 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed3_5 a b () f g h) = IPv6AddrPreferred a b 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_6 a b () e f g h) = IPv6AddrPreferred a b 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_7 a b () d e f g h) = IPv6AddrPreferred a b 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed4_3 a b c ()) = IPv6AddrPreferred a b c 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed4_4 a b c () h) = IPv6AddrPreferred a b c 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed4_5 a b c () g h) = IPv6AddrPreferred a b c 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed4_6 a b c () f g h) = IPv6AddrPreferred a b c 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed4_7 a b c () e f g h) = IPv6AddrPreferred a b c 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed5_4 a b c d ()) = IPv6AddrPreferred a b c d 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed5_5 a b c d () h) = IPv6AddrPreferred a b c d 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed5_6 a b c d () g h) = IPv6AddrPreferred a b c d 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed5_7 a b c d () f g h) = IPv6AddrPreferred a b c d 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed6_5 a b c d e ()) = IPv6AddrPreferred a b c d e 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed6_6 a b c d e () h) = IPv6AddrPreferred a b c d e 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed6_7 a b c d e () g h) = IPv6AddrPreferred a b c d e 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed7_6 a b c d e f ()) = IPv6AddrPreferred a b c d e f 0 0" | "ipv6addr_c2p (IPv6AddrCompressed7_7 a b c d e f () h) = IPv6AddrPreferred a b c d e f 0 h" | "ipv6addr_c2p (IPv6AddrCompressed8_7 a b c d e f g ()) = IPv6AddrPreferred a b c d e f g 0" definition ipv6_unparsed_compressed_to_preferred :: "((16 word) option) list \ ipv6addr_syntax option" where "ipv6_unparsed_compressed_to_preferred ls = ( if length (filter (\p. p = None) ls) \ 1 \ length (filter (\p. p \ None) ls) > 7 then None else let before_omission = map the (takeWhile (\x. x \ None) ls); after_omission = map the (drop 1 (dropWhile (\x. x \ None) ls)); num_omissions = 8 - (length before_omission + length after_omission); expanded = before_omission @ (replicate num_omissions 0) @ after_omission in case expanded of [a,b,c,d,e,f,g,h] \ Some (IPv6AddrPreferred a b c d e f g h) | _ \ None )" lemma "ipv6_unparsed_compressed_to_preferred [Some 0x2001, Some 0xDB8, None, Some 0x8, Some 0x800, Some 0x200C, Some 0x417A] = Some (IPv6AddrPreferred 0x2001 0xDB8 0 0 8 0x800 0x200C 0x417A)" by eval lemma "ipv6_unparsed_compressed_to_preferred [None] = Some (IPv6AddrPreferred 0 0 0 0 0 0 0 0)" by eval lemma "ipv6_unparsed_compressed_to_preferred [] = None" by eval lemma ipv6_unparsed_compressed_to_preferred_identity1: "ipv6_unparsed_compressed_to_preferred (ipv6addr_syntax_compressed_to_list ipv6compressed) = Some ipv6prferred \ ipv6addr_c2p ipv6compressed = ipv6prferred" by (cases ipv6compressed) (simp_all add: ipv6_unparsed_compressed_to_preferred_def numeral_eq_Suc) (*1s*) lemma ipv6_unparsed_compressed_to_preferred_identity2: "ipv6_unparsed_compressed_to_preferred ls = Some ipv6prferred \ (\ipv6compressed. parse_ipv6_address_compressed ls = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ipv6prferred)" apply(rule iffI) apply(subgoal_tac "parse_ipv6_address_compressed ls \ None") prefer 2 apply(subst RFC_4291_format) apply(simp add: ipv6_unparsed_compressed_to_preferred_def split: if_split_asm; fail) apply(simp) apply(erule exE, rename_tac ipv6compressed) apply(rule_tac x="ipv6compressed" in exI) apply(simp) apply(subgoal_tac "(ipv6addr_syntax_compressed_to_list ipv6compressed = ls)") prefer 2 using parse_ipv6_address_compressed_identity2 apply presburger using ipv6_unparsed_compressed_to_preferred_identity1 apply blast apply(erule exE, rename_tac ipv6compressed) apply(subgoal_tac "(ipv6addr_syntax_compressed_to_list ipv6compressed = ls)") prefer 2 using parse_ipv6_address_compressed_identity2 apply presburger using ipv6_unparsed_compressed_to_preferred_identity1 apply blast done subsection\IPv6 Pretty Printing (converting to compressed format)\ text_raw\ RFC5952: \begin{verbatim} 4. A Recommendation for IPv6 Text Representation A recommendation for a canonical text representation format of IPv6 addresses is presented in this section. The recommendation in this document is one that complies fully with [RFC4291], is implemented by various operating systems, and is human friendly. The recommendation in this section SHOULD be followed by systems when generating an address to be represented as text, but all implementations MUST accept and be able to handle any legitimate [RFC4291] format. It is advised that humans also follow these recommendations when spelling an address. 4.1. Handling Leading Zeros in a 16-Bit Field Leading zeros MUST be suppressed. For example, 2001:0db8::0001 is not acceptable and must be represented as 2001:db8::1. A single 16- bit 0000 field MUST be represented as 0. 4.2. "::" Usage 4.2.1. Shorten as Much as Possible The use of the symbol "::" MUST be used to its maximum capability. For example, 2001:db8:0:0:0:0:2:1 must be shortened to 2001:db8::2:1. Likewise, 2001:db8::0:1 is not acceptable, because the symbol "::" could have been used to produce a shorter representation 2001:db8::1. 4.2.2. Handling One 16-Bit 0 Field The symbol "::" MUST NOT be used to shorten just one 16-bit 0 field. For example, the representation 2001:db8:0:1:1:1:1:1 is correct, but 2001:db8::1:1:1:1:1 is not correct. 4.2.3. Choice in Placement of "::" When there is an alternative choice in the placement of a "::", the longest run of consecutive 16-bit 0 fields MUST be shortened (i.e., the sequence with three consecutive zero fields is shortened in 2001: 0:0:1:0:0:0:1). When the length of the consecutive 16-bit 0 fields are equal (i.e., 2001:db8:0:0:1:0:0:1), the first sequence of zero bits MUST be shortened. For example, 2001:db8::1:0:0:1 is correct representation. 4.3. Lowercase The characters "a", "b", "c", "d", "e", and "f" in an IPv6 address MUST be represented in lowercase. \end{verbatim} \ text\See @{file \IP_Address_toString.thy\} for examples and test cases.\ context begin private function goup_by_zeros :: "16 word list \ 16 word list list" where "goup_by_zeros [] = []" | "goup_by_zeros (x#xs) = ( if x = 0 then takeWhile (\x. x = 0) (x#xs) # (goup_by_zeros (dropWhile (\x. x = 0) xs)) else [x]#(goup_by_zeros xs))" by(pat_completeness, auto) termination goup_by_zeros apply(relation "measure (\xs. length xs)") apply(simp_all) by (simp add: le_imp_less_Suc length_dropWhile_le) private lemma "goup_by_zeros [0,1,2,3,0,0,0,0,3,4,0,0,0,2,0,0,2,0,3,0] = [[0], [1], [2], [3], [0, 0, 0, 0], [3], [4], [0, 0, 0], [2], [0, 0], [2], [0], [3], [0]]" by eval private lemma "concat (goup_by_zeros ls) = ls" by(induction ls rule:goup_by_zeros.induct) simp+ private lemma "[] \ set (goup_by_zeros ls)" by(induction ls rule:goup_by_zeros.induct) simp+ private primrec List_replace1 :: "'a \ 'a \ 'a list \ 'a list" where "List_replace1 _ _ [] = []" | "List_replace1 a b (x#xs) = (if a = x then b#xs else x#List_replace1 a b xs)" private lemma "List_replace1 a a ls = ls" by(induction ls) simp_all private lemma "a \ set ls \ List_replace1 a b ls = ls" by(induction ls) simp_all private lemma "a \ set ls \ b \ set (List_replace1 a b ls)" apply(induction ls) apply(simp) apply(simp) by blast private fun List_explode :: "'a list list \ ('a option) list" where "List_explode [] = []" | "List_explode ([]#xs) = None#List_explode xs" | "List_explode (xs1#xs2) = map Some xs1@List_explode xs2" private lemma "List_explode [[0::int], [2,3], [], [3,4]] = [Some 0, Some 2, Some 3, None, Some 3, Some 4]" by eval private lemma List_explode_def: "List_explode xss = concat (map (\xs. if xs = [] then [None] else map Some xs) xss)" by(induction xss rule: List_explode.induct) simp+ private lemma List_explode_no_empty: "[] \ set xss \ List_explode xss = map Some (concat xss)" by(induction xss rule: List_explode.induct) simp+ private lemma List_explode_replace1: "[] \ set xss \ foo \ set xss \ List_explode (List_replace1 foo [] xss) = map Some (concat (takeWhile (\xs. xs \ foo) xss)) @ [None] @ map Some (concat (tl (dropWhile (\xs. xs \ foo) xss)))" apply(induction xss rule: List_explode.induct) apply(simp; fail) apply(simp; fail) apply(simp) apply safe apply(simp_all add: List_explode_no_empty) done fun ipv6_preferred_to_compressed :: "ipv6addr_syntax \ ((16 word) option) list" where "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = ( let lss = goup_by_zeros [a,b,c,d,e,f,g,h]; max_zero_seq = foldr (\xs. max (length xs)) lss 0; shortened = if max_zero_seq > 1 then List_replace1 (replicate max_zero_seq 0) [] lss else lss in List_explode shortened )" declare ipv6_preferred_to_compressed.simps[simp del] private lemma foldr_max_length: "foldr (\xs. max (length xs)) lss n = fold max (map length lss) n" apply(subst List.foldr_fold) apply fastforce apply(induction lss arbitrary: n) apply(simp; fail) apply(simp) done private lemma List_explode_goup_by_zeros: "List_explode (goup_by_zeros xs) = map Some xs" apply(induction xs rule: goup_by_zeros.induct) apply(simp; fail) apply(simp) apply(safe) apply(simp) by (metis map_append takeWhile_dropWhile_id) private definition "max_zero_streak xs \ foldr (\xs. max (length xs)) (goup_by_zeros xs) 0" private lemma max_zero_streak_def2: "max_zero_streak xs = fold max (map length (goup_by_zeros xs)) 0" unfolding max_zero_streak_def by(simp add: foldr_max_length) private lemma ipv6_preferred_to_compressed_pull_out_if: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = ( if max_zero_streak [a,b,c,d,e,f,g,h] > 1 then List_explode (List_replace1 (replicate (max_zero_streak [a,b,c,d,e,f,g,h]) 0) [] (goup_by_zeros [a,b,c,d,e,f,g,h])) else map Some [a,b,c,d,e,f,g,h] )" by(simp add: ipv6_preferred_to_compressed.simps max_zero_streak_def List_explode_goup_by_zeros) private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0 0 0 0 0 0 0 0) = [None]" by eval private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0x2001 0xDB8 0 0 8 0x800 0x200C 0x417A) = [Some 0x2001, Some 0xDB8, None, Some 8, Some 0x800, Some 0x200C, Some 0x417A]" by eval private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0x2001 0xDB8 0 3 8 0x800 0x200C 0x417A) = [Some 0x2001, Some 0xDB8, Some 0, Some 3, Some 8, Some 0x800, Some 0x200C, Some 0x417A]" by eval (*the output should even conform to RFC5952, ...*) lemma ipv6_preferred_to_compressed_RFC_4291_format: "ipv6_preferred_to_compressed ip = as \ length (filter (\p. p = None) as) = 0 \ length as = 8 \ length (filter (\p. p = None) as) = 1 \ length (filter (\p. p \ None) as) \ 7" apply(cases ip) apply(simp add: ipv6_preferred_to_compressed_pull_out_if) apply(simp only: split: if_split_asm) subgoal for a b c d e f g h apply(rule disjI2) apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(auto simp add: max_zero_streak_def) (*1min*) subgoal apply(rule disjI1) apply(simp) by force done \ \Idea for the following proof:\ private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = None#xs \ xs = map Some (dropWhile (\x. x=0) [a,b,c,d,e,f,g,h])" apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*20s*) lemma ipv6_preferred_to_compressed: assumes "ipv6_unparsed_compressed_to_preferred (ipv6_preferred_to_compressed ip) = Some ip'" shows "ip = ip'" proof - from assms have 1: "\ipv6compressed. parse_ipv6_address_compressed (ipv6_preferred_to_compressed ip) = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ip'" using ipv6_unparsed_compressed_to_preferred_identity2 by simp obtain a b c d e f g h where ip: "ip = IPv6AddrPreferred a b c d e f g h" by(cases ip) have ipv6_preferred_to_compressed_None1: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = None#xs \ (map Some (dropWhile (\x. x=0) [a,b,c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a b c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None2: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#None#xs \ (map Some (dropWhile (\x. x=0) [b,c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None3: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#None#xs \ (map Some (dropWhile (\x. x=0) [c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None4: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#None#xs \ (map Some (dropWhile (\x. x=0) [d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None5: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#None#xs \ (map Some (dropWhile (\x. x=0) [e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None6: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#None#xs \ (map Some (dropWhile (\x. x=0) [f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None7: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#(Some f')#None#xs \ (map Some (dropWhile (\x. x=0) [g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f' g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' f' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None8: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#(Some f')#(Some g')#None#xs \ (map Some (dropWhile (\x. x=0) [h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f' g' h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' f' g' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have 2: "parse_ipv6_address_compressed (ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h)) = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ip' \ IPv6AddrPreferred a b c d e f g h = ip'" for ipv6compressed apply(erule parse_ipv6_address_compressed_someE) apply(simp_all) apply(erule ipv6_preferred_to_compressed_None1, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None2, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None3, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None4, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None5, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None6, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None7, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None8, simp split: if_split_asm) done from 1 2 ip show ?thesis by(elim exE conjE, simp) qed end end diff --git a/thys/IP_Addresses/WordInterval.thy b/thys/IP_Addresses/WordInterval.thy --- a/thys/IP_Addresses/WordInterval.thy +++ b/thys/IP_Addresses/WordInterval.thy @@ -1,777 +1,777 @@ (* Title: WordInterval.thy Authors: Julius Michaelis, Cornelius Diekmann *) theory WordInterval imports Main "Word_Lib.Word_Lemmas" "Word_Lib.Next_and_Prev" begin section\WordInterval: Executable datatype for Machine Word Sets\ text\Stores ranges of machine words as interval. This has been proven quite efficient for IP Addresses.\ (*NOTE: All algorithms here use a straight-forward implementation. There is a lot of room for improving the computation complexity, for example by making the WordInterval a balanced, sorted tree.*) subsection\Syntax\ context notes [[typedef_overloaded]] begin datatype ('a::len) wordinterval = WordInterval "('a::len) word" \ \start (inclusive)\ "('a::len) word" \ \end (inclusive)\ | RangeUnion "'a wordinterval" "'a wordinterval" end subsection\Semantics\ fun wordinterval_to_set :: "'a::len wordinterval \ ('a::len word) set" where "wordinterval_to_set (WordInterval start end) = {start .. end}" | "wordinterval_to_set (RangeUnion r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" (*Note: The runtime of all the operations could be improved, for example by keeping the tree sorted and balanced.*) subsection\Basic operations\ text\\\\\ fun wordinterval_element :: "'a::len word \ 'a::len wordinterval \ bool" where "wordinterval_element el (WordInterval s e) \ s \ el \ el \ e" | "wordinterval_element el (RangeUnion r1 r2) \ wordinterval_element el r1 \ wordinterval_element el r2" lemma wordinterval_element_set_eq[simp]: "wordinterval_element el rg = (el \ wordinterval_to_set rg)" by(induction rg rule: wordinterval_element.induct) simp_all definition wordinterval_union :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_union r1 r2 = RangeUnion r1 r2" lemma wordinterval_union_set_eq[simp]: "wordinterval_to_set (wordinterval_union r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" unfolding wordinterval_union_def by simp fun wordinterval_empty :: "'a::len wordinterval \ bool" where "wordinterval_empty (WordInterval s e) \ e < s" | "wordinterval_empty (RangeUnion r1 r2) \ wordinterval_empty r1 \ wordinterval_empty r2" lemma wordinterval_empty_set_eq[simp]: "wordinterval_empty r \ wordinterval_to_set r = {}" by(induction r) auto definition Empty_WordInterval :: "'a::len wordinterval" where "Empty_WordInterval \ WordInterval 1 0" lemma wordinterval_empty_Empty_WordInterval: "wordinterval_empty Empty_WordInterval" by(simp add: Empty_WordInterval_def) lemma Empty_WordInterval_set_eq[simp]: "wordinterval_to_set Empty_WordInterval = {}" by(simp add: Empty_WordInterval_def) subsection\WordInterval and Lists\ text\A list of \(start, end)\ tuples.\ text\wordinterval to list\ fun wi2l :: "'a::len wordinterval \ ('a::len word \ 'a::len word) list" where "wi2l (RangeUnion r1 r2) = wi2l r1 @ wi2l r2" | "wi2l (WordInterval s e) = (if e < s then [] else [(s,e)])" text\list to wordinterval\ fun l2wi :: "('a::len word \ 'a word) list \ 'a wordinterval" where "l2wi [] = Empty_WordInterval" | "l2wi [(s,e)] = (WordInterval s e)" | "l2wi ((s,e)#rs) = (RangeUnion (WordInterval s e) (l2wi rs))" lemma l2wi_append: "wordinterval_to_set (l2wi (l1@l2)) = wordinterval_to_set (l2wi l1) \ wordinterval_to_set (l2wi l2)" proof(induction l1 arbitrary: l2 rule:l2wi.induct) case 1 thus ?case by simp next case (2 s e l2) thus ?case by (cases l2) simp_all next case 3 thus ?case by force qed lemma l2wi_wi2l[simp]: "wordinterval_to_set (l2wi (wi2l r)) = wordinterval_to_set r" by(induction r) (simp_all add: l2wi_append) lemma l2wi: "wordinterval_to_set (l2wi l) = (\ (i,j) \ set l. {i .. j})" by(induction l rule: l2wi.induct, simp_all) lemma wi2l: "(\(i,j)\set (wi2l r). {i .. j}) = wordinterval_to_set r" by(induction r rule: wi2l.induct, simp_all) lemma l2wi_remdups[simp]: "wordinterval_to_set (l2wi (remdups ls)) = wordinterval_to_set (l2wi ls)" by(simp add: l2wi) lemma wi2l_empty[simp]: "wi2l Empty_WordInterval = []" unfolding Empty_WordInterval_def by simp subsection\Optimizing and minimizing @{typ "('a::len) wordinterval"}s\ text\Removing empty intervals\ context begin fun wordinterval_optimize_empty :: "'a::len wordinterval \ 'a wordinterval" where "wordinterval_optimize_empty (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty r1o then r2o else if wordinterval_empty r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty r = r" lemma wordinterval_optimize_empty_set_eq[simp]: "wordinterval_to_set (wordinterval_optimize_empty r) = wordinterval_to_set r" by(induction r) (simp_all add: Let_def) lemma wordinterval_optimize_empty_double: "wordinterval_optimize_empty (wordinterval_optimize_empty r) = wordinterval_optimize_empty r" by(induction r) (simp_all add: Let_def) private fun wordinterval_empty_shallow :: "'a::len wordinterval \ bool" where "wordinterval_empty_shallow (WordInterval s e) \ e < s" | "wordinterval_empty_shallow (RangeUnion _ _) \ False" private lemma helper_optimize_shallow: "wordinterval_empty_shallow (wordinterval_optimize_empty r) = wordinterval_empty (wordinterval_optimize_empty r)" by(induction r) fastforce+ private fun wordinterval_optimize_empty2 where "wordinterval_optimize_empty2 (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty_shallow r1o then r2o else if wordinterval_empty_shallow r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty2 r = r" lemma wordinterval_optimize_empty_code[code_unfold]: "wordinterval_optimize_empty = wordinterval_optimize_empty2" by (subst fun_eq_iff, clarify, rename_tac r, induct_tac r) (unfold wordinterval_optimize_empty.simps wordinterval_optimize_empty2.simps Let_def helper_optimize_shallow, simp_all) end text\Merging overlapping intervals\ context begin private definition disjoint :: "'a set \ 'a set \ bool" where "disjoint A B \ A \ B = {}" private primrec interval_of :: "('a::len) word \ 'a word \ 'a word set" where "interval_of (s,e) = {s .. e}" declare interval_of.simps[simp del] private definition disjoint_intervals :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "disjoint_intervals A B \ disjoint (interval_of A) (interval_of B)" private definition not_disjoint_intervals :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "not_disjoint_intervals A B \ \ disjoint (interval_of A) (interval_of B)" private lemma [code]: "not_disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s \ e' \ s' \ e \ s \ e \ s' \ e')" apply(cases A, cases B) apply(simp add: not_disjoint_intervals_def interval_of.simps disjoint_def) done private lemma [code]: "disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s > e' \ s' > e \ s > e \ s' > e')" apply(cases A, cases B) apply(simp add: disjoint_intervals_def interval_of.simps disjoint_def) by fastforce text\BEGIN merging overlapping intervals\ (*result has no empty intervals and all are disjoint. merging things such as [1,7] [8,10] would still be possible*) private fun merge_overlap :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_overlap s [] = [s]" | "merge_overlap (s,e) ((s',e')#ss) = ( if not_disjoint_intervals (s,e) (s',e') then (min s s', max e e')#ss else (s',e')#merge_overlap (s,e) ss)" private lemma not_disjoint_union: fixes s :: "('a::len) word" shows "\ disjoint {s..e} {s'..e'} \ {s..e} \ {s'..e'} = {min s s' .. max e e'}" by(auto simp add: disjoint_def min_def max_def) private lemma disjoint_subset: "disjoint A B \ A \ B \ C \ A \ C" unfolding disjoint_def by blast private lemma merge_overlap_helper1: "interval_of A \ (\s \ set ss. interval_of s) \ (\s \ set (merge_overlap A ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(drule_tac C="(\x\set xs. interval_of x)" in disjoint_subset) apply(simp_all) done private lemma merge_overlap_helper2: "\s'\set ss. \ disjoint (interval_of A) (interval_of s') \ interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_overlap A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(simp) by blast private lemma merge_overlap_length: "\s' \ set ss. \ disjoint (interval_of A) (interval_of s') \ length (merge_overlap A ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) done lemma "merge_overlap (1:: 16 word,2) [(1, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(2, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(3, 7)] = [(3, 7), (1,2)]" by eval private function listwordinterval_compress :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_compress [] = []" | "listwordinterval_compress (s#ss) = ( if \s' \ set ss. disjoint_intervals s s' then s#listwordinterval_compress ss else listwordinterval_compress (merge_overlap s ss))" by(pat_completeness, auto) private termination listwordinterval_compress apply (relation "measure length") apply(rule wf_measure) apply(simp) using disjoint_intervals_def merge_overlap_length by fastforce private lemma listwordinterval_compress: "(\s \ set (listwordinterval_compress ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_compress.induct) apply(simp) apply(simp) apply(intro impI) apply(simp add: disjoint_intervals_def) apply(drule merge_overlap_helper2) apply(simp) done lemma "listwordinterval_compress [(1::32 word,3), (8,10), (2,5), (3,7)] = [(8, 10), (1, 7)]" by eval private lemma A_in_listwordinterval_compress: "A \ set (listwordinterval_compress ss) \ interval_of A \ (\s \ set ss. interval_of s)" using listwordinterval_compress by blast private lemma listwordinterval_compress_disjoint: "A \ set (listwordinterval_compress ss) \ B \ set (listwordinterval_compress ss) \ A \ B \ disjoint (interval_of A) (interval_of B)" apply(induction ss arbitrary: rule: listwordinterval_compress.induct) apply(simp) apply(simp split: if_split_asm) apply(elim disjE) apply(simp_all) apply(simp_all add: disjoint_intervals_def disjoint_def) apply(blast dest: A_in_listwordinterval_compress)+ done text\END merging overlapping intervals\ text\BEGIN merging adjacent intervals\ private fun merge_adjacent :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_adjacent s [] = [s]" | "merge_adjacent (s,e) ((s',e')#ss) = ( if s \e \ s' \ e' \ word_next e = s' then (s, e')#ss else if s \e \ s' \ e' \ word_next e' = s then (s', e)#ss else (s',e')#merge_adjacent (s,e) ss)" private lemma merge_adjacent_helper: "interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_adjacent A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: interval_of.simps) apply(intro impI conjI) apply (metis Un_assoc word_adjacent_union) apply(elim conjE) apply(drule(2) word_adjacent_union) subgoal by (blast) subgoal by (metis word_adjacent_union Un_assoc) by blast private lemma merge_adjacent_length: "\(s', e')\set ss. s \ e \ s' \ e' \ (word_next e = s' \ word_next e' = s) \ length (merge_adjacent (s,e) ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(case_tac x) apply(simp add: ) by blast private function listwordinterval_adjacent :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_adjacent [] = []" | "listwordinterval_adjacent ((s,e)#ss) = ( if \(s',e') \ set ss. \ (s \e \ s' \ e' \ (word_next e = s' \ word_next e' = s)) then (s,e)#listwordinterval_adjacent ss else listwordinterval_adjacent (merge_adjacent (s,e) ss))" by(pat_completeness, auto) private termination listwordinterval_adjacent apply (relation "measure length") apply(rule wf_measure) apply(simp) apply(simp) using merge_adjacent_length by fastforce private lemma listwordinterval_adjacent: "(\s \ set (listwordinterval_adjacent ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_adjacent.induct) apply(simp) apply(simp add: merge_adjacent_helper) done lemma "listwordinterval_adjacent [(1::16 word, 3), (5, 10), (10,10), (4,4)] = [(10, 10), (1, 10)]" by eval text\END merging adjacent intervals\ definition wordinterval_compress :: "('a::len) wordinterval \ 'a wordinterval" where "wordinterval_compress r \ l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))" text\Correctness: Compression preserves semantics\ lemma wordinterval_compress: "wordinterval_to_set (wordinterval_compress r) = wordinterval_to_set r" unfolding wordinterval_compress_def proof - have interval_of': "interval_of s = (case s of (s,e) \ {s .. e})" for s by (cases s) (simp add: interval_of.simps) have "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = (\x\set (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))). interval_of x)" by (force simp: interval_of' l2wi) also have "\ = (\s\set (wi2l (wordinterval_optimize_empty r)). interval_of s)" by(simp add: listwordinterval_compress listwordinterval_adjacent) also have "\ = (\(i, j)\set (wi2l (wordinterval_optimize_empty r)). {i..j})" by(simp add: interval_of') also have "\ = wordinterval_to_set r" by(simp add: wi2l) finally show "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = wordinterval_to_set r" . qed end text\Example\ lemma "(wi2l \ (wordinterval_compress :: 32 wordinterval \ 32 wordinterval) \ l2wi) [(70, 80001), (0,0), (150, 8000), (1,3), (42,41), (3,7), (56, 200), (8,10)] = [(56, 80001), (0, 10)]" by eval lemma "wordinterval_compress (RangeUnion (RangeUnion (WordInterval (1::32 word) 5) (WordInterval 8 10)) (WordInterval 3 7)) = WordInterval 1 10" by eval subsection\Further operations\ text\\\\\ definition wordinterval_Union :: "('a::len) wordinterval list \ 'a wordinterval" where "wordinterval_Union ws = wordinterval_compress (foldr wordinterval_union ws Empty_WordInterval)" lemma wordinterval_Union: "wordinterval_to_set (wordinterval_Union ws) = (\ w \ (set ws). wordinterval_to_set w)" by(induction ws) (simp_all add: wordinterval_compress wordinterval_Union_def) context begin private fun wordinterval_setminus' :: "'a::len wordinterval \ 'a wordinterval \ 'a wordinterval" where "wordinterval_setminus' (WordInterval s e) (WordInterval ms me) = ( if s > e \ ms > me then WordInterval s e else if me \ e then WordInterval (if ms = 0 then 1 else s) (min e (word_prev ms)) else if ms \ s then - WordInterval (max s (word_next me)) (if me = max_word then 0 else e) + WordInterval (max s (word_next me)) (if me = - 1 then 0 else e) else RangeUnion (WordInterval (if ms = 0 then 1 else s) (word_prev ms)) - (WordInterval (word_next me) (if me = max_word then 0 else e)) + (WordInterval (word_next me) (if me = - 1 then 0 else e)) )" | "wordinterval_setminus' (RangeUnion r1 r2) t = RangeUnion (wordinterval_setminus' r1 t) (wordinterval_setminus' r2 t)"| "wordinterval_setminus' t (RangeUnion r1 r2) = wordinterval_setminus' (wordinterval_setminus' t r1) r2" private lemma wordinterval_setminus'_rr_set_eq: "wordinterval_to_set(wordinterval_setminus' (WordInterval s e) (WordInterval ms me)) = wordinterval_to_set (WordInterval s e) - wordinterval_to_set (WordInterval ms me)" apply(simp only: wordinterval_setminus'.simps) apply(case_tac "e < s") apply simp apply(case_tac "me < ms") apply simp apply(case_tac [!] "e \ me") apply(case_tac [!] "ms = 0") apply(case_tac [!] "ms \ s") - apply(case_tac [!] "me = max_word") + apply(case_tac [!] "me = - 1") apply(simp_all add: word_next_unfold word_prev_unfold min_def max_def) apply(safe) apply(auto) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) done private lemma wordinterval_setminus'_set_eq: "wordinterval_to_set (wordinterval_setminus' r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" apply(induction rule: wordinterval_setminus'.induct) using wordinterval_setminus'_rr_set_eq apply blast apply auto done lemma wordinterval_setminus'_empty_struct: "wordinterval_empty r2 \ wordinterval_setminus' r1 r2 = r1" by(induction r1 r2 rule: wordinterval_setminus'.induct) auto definition wordinterval_setminus :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_setminus r1 r2 = wordinterval_compress (wordinterval_setminus' r1 r2)" lemma wordinterval_setminus_set_eq[simp]: "wordinterval_to_set (wordinterval_setminus r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" by(simp add: wordinterval_setminus_def wordinterval_compress wordinterval_setminus'_set_eq) end definition wordinterval_UNIV :: "'a::len wordinterval" where - "wordinterval_UNIV \ WordInterval 0 max_word" + "wordinterval_UNIV \ WordInterval 0 (- 1)" lemma wordinterval_UNIV_set_eq[simp]: "wordinterval_to_set wordinterval_UNIV = UNIV" unfolding wordinterval_UNIV_def using max_word_max by fastforce fun wordinterval_invert :: "'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_invert r = wordinterval_setminus wordinterval_UNIV r" lemma wordinterval_invert_set_eq[simp]: "wordinterval_to_set (wordinterval_invert r) = UNIV - wordinterval_to_set r" by(auto) lemma wordinterval_invert_UNIV_empty: "wordinterval_empty (wordinterval_invert wordinterval_UNIV)" by simp -lemma wi2l_univ[simp]: "wi2l wordinterval_UNIV = [(0, max_word)]" +lemma wi2l_univ[simp]: "wi2l wordinterval_UNIV = [(0, - 1)]" unfolding wordinterval_UNIV_def by simp text\\\\\ context begin private lemma "{(s::nat) .. e} \ {s' .. e'} = {} \ s > e' \ s' > e \ s > e \ s' > e'" by simp linarith private fun wordinterval_intersection' :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection' (WordInterval s e) (WordInterval s' e') = ( if s > e \ s' > e' \ s > e' \ s' > e \ s > e \ s' > e' then Empty_WordInterval else WordInterval (max s s') (min e e') )" | "wordinterval_intersection' (RangeUnion r1 r2) t = RangeUnion (wordinterval_intersection' r1 t) (wordinterval_intersection' r2 t)"| "wordinterval_intersection' t (RangeUnion r1 r2) = RangeUnion (wordinterval_intersection' t r1) (wordinterval_intersection' t r2)" private lemma wordinterval_intersection'_set_eq: "wordinterval_to_set (wordinterval_intersection' r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(induction r1 r2 rule: wordinterval_intersection'.induct) (auto) lemma "wordinterval_intersection' (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = RangeUnion (RangeUnion (WordInterval 1 3) (WordInterval 1 0)) (WordInterval 1 3)" by eval definition wordinterval_intersection :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection r1 r2 \ wordinterval_compress (wordinterval_intersection' r1 r2)" lemma wordinterval_intersection_set_eq[simp]: "wordinterval_to_set (wordinterval_intersection r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(simp add: wordinterval_intersection_def wordinterval_compress wordinterval_intersection'_set_eq) lemma "wordinterval_intersection (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = WordInterval 1 3" by eval end definition wordinterval_subset :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_subset r1 r2 \ wordinterval_empty (wordinterval_setminus r1 r2)" lemma wordinterval_subset_set_eq[simp]: "wordinterval_subset r1 r2 = (wordinterval_to_set r1 \ wordinterval_to_set r2)" unfolding wordinterval_subset_def by simp definition wordinterval_eq :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_eq r1 r2 = (wordinterval_subset r1 r2 \ wordinterval_subset r2 r1)" lemma wordinterval_eq_set_eq: "wordinterval_eq r1 r2 \ wordinterval_to_set r1 = wordinterval_to_set r2" unfolding wordinterval_eq_def by auto thm iffD1[OF wordinterval_eq_set_eq] (*declare iffD1[OF wordinterval_eq_set_eq, simp]*) lemma wordinterval_eq_comm: "wordinterval_eq r1 r2 \ wordinterval_eq r2 r1" unfolding wordinterval_eq_def by fast lemma wordinterval_to_set_alt: "wordinterval_to_set r = {x. wordinterval_element x r}" unfolding wordinterval_element_set_eq by blast lemma wordinterval_un_empty: "wordinterval_empty r1 \ wordinterval_eq (wordinterval_union r1 r2) r2" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_un_emty_b: "wordinterval_empty r2 \ wordinterval_eq (wordinterval_union r1 r2) r1" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_Diff_triv: "wordinterval_empty (wordinterval_intersection a b) \ wordinterval_eq (wordinterval_setminus a b) a" unfolding wordinterval_eq_set_eq by simp blast text\A size of the datatype, does not correspond to the cardinality of the corresponding set\ fun wordinterval_size :: "('a::len) wordinterval \ nat" where "wordinterval_size (RangeUnion a b) = wordinterval_size a + wordinterval_size b" | "wordinterval_size (WordInterval s e) = (if s \ e then 1 else 0)" lemma wordinterval_size_length: "wordinterval_size r = length (wi2l r)" by(induction r) (auto) lemma Ex_wordinterval_nonempty: "\x::('a::len wordinterval). y \ wordinterval_to_set x" proof show "y \ wordinterval_to_set wordinterval_UNIV" by simp qed lemma wordinterval_eq_reflp: "reflp wordinterval_eq" apply(rule reflpI) by(simp only: wordinterval_eq_set_eq) lemma wordintervalt_eq_symp: "symp wordinterval_eq" apply(rule sympI) by(simp add: wordinterval_eq_comm) lemma wordinterval_eq_transp: "transp wordinterval_eq" apply(rule transpI) by(simp only: wordinterval_eq_set_eq) lemma wordinterval_eq_equivp: "equivp wordinterval_eq" by (auto intro: equivpI wordinterval_eq_reflp wordintervalt_eq_symp wordinterval_eq_transp) text\The smallest element in the interval\ definition is_lowest_element :: "'a::ord \ 'a set \ bool" where "is_lowest_element x S = (x \ S \ (\y\S. y \ x \ y = x))" lemma fixes x :: "'a :: complete_lattice" assumes "x \ S" shows " x = Inf S \ is_lowest_element x S" using assms apply(simp add: is_lowest_element_def) by (simp add: Inf_lower eq_iff) lemma fixes x :: "'a :: linorder" assumes "finite S" and "x \ S" shows "is_lowest_element x S \ x = Min S" apply(rule) subgoal apply(simp add: is_lowest_element_def) apply(subst Min_eqI[symmetric]) using assms by(auto) by (metis Min.coboundedI assms(1) assms(2) dual_order.antisym is_lowest_element_def) text\Smallest element in the interval\ fun wordinterval_lowest_element :: "'a::len wordinterval \ 'a word option" where "wordinterval_lowest_element (WordInterval s e) = (if s \ e then Some s else None)" | "wordinterval_lowest_element (RangeUnion A B) = (case (wordinterval_lowest_element A, wordinterval_lowest_element B) of (Some a, Some b) \ Some (if a < b then a else b) | (None, Some b) \ Some b | (Some a, None) \ Some a | (None, None) \ None)" lemma wordinterval_lowest_none_empty: "wordinterval_lowest_element r = None \ wordinterval_empty r" proof(induction r) case WordInterval thus ?case by simp next case RangeUnion thus ?case by fastforce qed lemma wordinterval_lowest_element_correct_A: "wordinterval_lowest_element r = Some x \ is_lowest_element x (wordinterval_to_set r)" unfolding is_lowest_element_def apply(induction r arbitrary: x rule: wordinterval_lowest_element.induct) apply(rename_tac rs re x, case_tac "rs \ re", auto)[1] apply(subst(asm) wordinterval_lowest_element.simps(2)) apply(rename_tac A B x) apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(simp_all add: wordinterval_lowest_none_empty)[3] apply fastforce done lemma wordinterval_lowest_element_set_eq: assumes "\ wordinterval_empty r" shows "(wordinterval_lowest_element r = Some x) = (is_lowest_element x (wordinterval_to_set r))" (*unfolding is_lowest_element_def*) proof(rule iffI) assume "wordinterval_lowest_element r = Some x" thus "is_lowest_element x (wordinterval_to_set r)" using wordinterval_lowest_element_correct_A wordinterval_lowest_none_empty by simp next assume "is_lowest_element x (wordinterval_to_set r)" with assms show "(wordinterval_lowest_element r = Some x)" proof(induction r arbitrary: x rule: wordinterval_lowest_element.induct) case 1 thus ?case by(simp add: is_lowest_element_def) next case (2 A B x) have is_lowest_RangeUnion: "is_lowest_element x (wordinterval_to_set A \ wordinterval_to_set B) \ is_lowest_element x (wordinterval_to_set A) \ is_lowest_element x (wordinterval_to_set B)" by(simp add: is_lowest_element_def) (*why \ A B?*) have wordinterval_lowest_element_RangeUnion: "\a b A B. wordinterval_lowest_element A = Some a \ wordinterval_lowest_element B = Some b \ wordinterval_lowest_element (RangeUnion A B) = Some (min a b)" by(auto dest!: wordinterval_lowest_element_correct_A simp add: is_lowest_element_def min_def) from 2 show ?case apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(auto simp add: is_lowest_element_def)[3] apply(subgoal_tac "\ wordinterval_empty A \ \ wordinterval_empty B") prefer 2 using arg_cong[where f = Not, OF wordinterval_lowest_none_empty] apply blast apply(drule(1) wordinterval_lowest_element_RangeUnion) apply(simp split: option.split_asm add: min_def) apply(drule is_lowest_RangeUnion) apply(elim disjE) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) using wordinterval_lowest_element_correct_A[simplified is_lowest_element_def] by (metis Un_iff not_le) qed qed text\Cardinality approximation for @{typ "('a::len) wordinterval"}s\ context begin lemma card_atLeastAtMost_word: fixes s::"('a::len) word" shows "card {s..e} = Suc (unat e) - (unat s)" apply(cases "s > e") apply(simp) apply(subst(asm) Word.word_less_nat_alt) apply simp apply(subst upto_enum_set_conv2[symmetric]) apply(subst List.card_set) apply(simp add: remdups_enum_upto) done fun wordinterval_card :: "('a::len) wordinterval \ nat" where "wordinterval_card (WordInterval s e) = Suc (unat e) - (unat s)" | "wordinterval_card (RangeUnion a b) = wordinterval_card a + wordinterval_card b" lemma wordinterval_card: "wordinterval_card r \ card (wordinterval_to_set r)" proof(induction r) case WordInterval thus ?case by (simp add: card_atLeastAtMost_word) next case (RangeUnion r1 r2) have "card (wordinterval_to_set r1 \ wordinterval_to_set r2) \ card (wordinterval_to_set r1) + card (wordinterval_to_set r2)" using Finite_Set.card_Un_le by blast with RangeUnion show ?case by(simp) qed text\With @{thm wordinterval_compress} it should be possible to get the exact cardinality\ end end diff --git a/thys/Iptables_Semantics/Common/Word_Upto.thy b/thys/Iptables_Semantics/Common/Word_Upto.thy --- a/thys/Iptables_Semantics/Common/Word_Upto.thy +++ b/thys/Iptables_Semantics/Common/Word_Upto.thy @@ -1,206 +1,206 @@ section\Word Upto\ theory Word_Upto imports Main IP_Addresses.Hs_Compat Word_Lib.Word_Lemmas begin text\Enumerate a range of machine words.\ text\enumerate from the back (inefficient)\ function word_upto :: "'a word \ 'a word \ ('a::len) word list" where "word_upto a b = (if a = b then [a] else word_upto a (b - 1) @ [b])" by pat_completeness auto (*by the way: does not terminate practically if b < a; will terminate after it reaches the word wrap-around!*) termination word_upto apply(relation "measure (unat \ uncurry (-) \ prod.swap)") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply(simp add: diff_right_commute; fail) apply(rule measure_unat) apply auto done declare word_upto.simps[simp del] text\enumerate from the front (more inefficient)\ function word_upto' :: "'a word \ 'a word \ ('a::len) word list" where "word_upto' a b = (if a = b then [a] else a # word_upto' (a + 1) b)" by pat_completeness auto termination word_upto' apply(relation "measure (\ (a, b). unat (b - a))") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply (simp add: diff_diff_add; fail) apply(rule measure_unat) apply auto done declare word_upto'.simps[simp del] lemma word_upto_cons_front[code]: "word_upto a b = word_upto' a b" proof(induction a b rule:word_upto'.induct) case (1 a b) have hlp1: "a \ b \ a # word_upto (a + 1) b = word_upto a b" apply(induction a b rule:word_upto.induct) apply simp apply(subst(1) word_upto.simps) apply(simp) apply safe apply(subst(1) word_upto.simps) apply (simp) apply(subst(1) word_upto.simps) apply (simp; fail) apply(case_tac "a \ b - 1") apply(simp) apply (metis Cons_eq_appendI word_upto.simps) apply(simp) done from 1[symmetric] show ?case apply(cases "a = b") subgoal apply(subst word_upto.simps) apply(subst word_upto'.simps) by(simp) apply(subst word_upto'.simps) by(simp add: hlp1) qed (* Most of the lemmas I show about word_upto hold without a \ b, but I don't need that right now and it's giving me a headache *) lemma word_upto_set_eq: "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" proof show "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst(asm) word_upto.simps) apply(simp; fail) apply(subst(asm) word_upto.simps) apply(simp) apply(erule disjE) apply(simp; fail) proof(goal_cases) case (1 a b) from 1(2-3) have "b \ 0" by force from 1(2,3) have "a \ b - 1" by (simp add: word_le_minus_one_leq) from 1(1)[OF this 1(4)] show ?case by (metis dual_order.trans 1(2,3) less_imp_le measure_unat word_le_0_iff word_le_nat_alt) qed next show "a \ x \ x \ b \ x \ set (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(simp) apply(case_tac "x = b") apply(simp;fail) proof(goal_cases) case (1 a b) from 1(2-4) have "b \ 0" by force from 1(2,4) have "x \ b - 1" using le_step_down_word by auto from 1(1) this show ?case by simp qed qed lemma word_upto_distinct_hlp: "a \ b \ a \ b \ b \ set (word_upto a (b - 1))" apply(rule ccontr, unfold not_not) apply(subgoal_tac "a \ b - 1") apply(drule iffD1[OF word_upto_set_eq[of a "b -1" b]]) apply(simp add: word_upto.simps) apply(subgoal_tac "b \ 0") apply(meson leD measure_unat word_le_nat_alt) apply(blast intro: iffD1[OF word_le_0_iff]) using le_step_down_word apply blast done lemma distinct_word_upto: "a \ b \ distinct (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(case_tac "a \ b - 1") apply(simp) apply(rule word_upto_distinct_hlp; simp) apply(simp) apply(rule ccontr) apply (simp add: not_le antisym word_minus_one_le_leq) done lemma word_upto_eq_upto: "s \ e \ e \ unat (max_word :: 'l word) \ word_upto ((of_nat :: nat \ ('l :: len) word) s) (of_nat e) = map of_nat (upt s (Suc e))" proof(induction e) let ?mwon = "of_nat :: nat \ 'l word" let ?mmw = "max_word :: 'l word" case (Suc e) show ?case proof(cases "?mwon s = ?mwon (Suc e)") case True have "s = Suc e" using le_unat_uoi Suc.prems True by metis with True show ?thesis by(subst word_upto.simps) (simp) next case False hence le: "s \ e" using le_SucE Suc.prems by blast have lm: "e \ unat ?mmw" using Suc.prems by simp have sucm: "(of_nat :: nat \ ('l :: len) word) (Suc e) - 1 = of_nat e" using Suc.prems(2) by simp note mIH = Suc.IH[OF le lm] show ?thesis by(subst word_upto.simps) (simp add: False[simplified] Suc.prems mIH sucm) qed qed(simp add: word_upto.simps) lemma word_upto_alt: "(a :: ('l :: len) word) \ b \ word_upto a b = map of_nat (upt (unat a) (Suc (unat b)))" proof - - let ?mmw = "max_word :: 'l word" + let ?mmw = "- 1 :: 'l word" assume le: "a \ b" hence nle: "unat a \ unat b" by(unat_arith) have lem: "unat b \ unat ?mmw" by (simp add: word_unat_less_le) note word_upto_eq_upto[OF nle lem, unfolded word_unat.Rep_inverse] thus "word_upto a b = map of_nat [unat a.. b then map of_nat (upt (unat a) (Suc (unat b))) else word_upto a b)" using word_upto_alt by metis lemma sorted_word_upto: fixes a b :: "('l :: len) word" assumes "a \ b" shows "sorted (word_upto a b)" proof - define m and n where \m = unat a\ and \n = Suc (unat b)\ moreover have \sorted (map of_nat [m.. apply (simp add: sorted_map) apply (rule sorted_wrt_mono_rel [of _ \(\)\]) apply simp_all apply (simp add: le_unat_uoi less_Suc_eq_le n_def word_of_nat_le) apply transfer apply simp apply (subst take_bit_int_eq_self) apply (simp_all add: le_less_trans) apply (metis le_unat_uoi of_int_of_nat_eq of_nat_mono uint_word_of_int_eq unat_eq_nat_uint unsigned_of_int) done ultimately have \sorted (map of_nat [unat a.. by simp with assms show ?thesis by (simp only: word_upto_alt) qed end diff --git a/thys/Iptables_Semantics/Primitive_Matchers/Common_Primitive_Matcher.thy b/thys/Iptables_Semantics/Primitive_Matchers/Common_Primitive_Matcher.thy --- a/thys/Iptables_Semantics/Primitive_Matchers/Common_Primitive_Matcher.thy +++ b/thys/Iptables_Semantics/Primitive_Matchers/Common_Primitive_Matcher.thy @@ -1,237 +1,237 @@ theory Common_Primitive_Matcher imports Common_Primitive_Matcher_Generic begin subsection\Primitive Matchers: IP Port Iface Matcher\ (*IPv4 matcher*) fun common_matcher :: "('i::len common_primitive, ('i, 'a) tagged_packet_scheme) exact_match_tac" where "common_matcher (IIface i) p = bool_to_ternary (match_iface i (p_iiface p))" | "common_matcher (OIface i) p = bool_to_ternary (match_iface i (p_oiface p))" | "common_matcher (Src ip) p = bool_to_ternary (p_src p \ ipt_iprange_to_set ip)" | "common_matcher (Dst ip) p = bool_to_ternary (p_dst p \ ipt_iprange_to_set ip)" | "common_matcher (Prot proto) p = bool_to_ternary (match_proto proto (p_proto p))" | "common_matcher (Src_Ports (L4Ports proto ps)) p = bool_to_ternary (proto = p_proto p \ p_sport p \ ports_to_set ps)" | "common_matcher (Dst_Ports (L4Ports proto ps)) p = bool_to_ternary (proto = p_proto p \ p_dport p \ ports_to_set ps)" | "common_matcher (MultiportPorts (L4Ports proto ps)) p = bool_to_ternary (proto = p_proto p \ (p_sport p \ ports_to_set ps \ p_dport p \ ports_to_set ps))" | "common_matcher (L4_Flags flags) p = bool_to_ternary (match_tcp_flags flags (p_tcp_flags p))" | "common_matcher (CT_State S) p = bool_to_ternary (match_ctstate S (p_tag_ctstate p))" | "common_matcher (Extra _) p = TernaryUnknown" lemma packet_independent_\_unknown_common_matcher: "packet_independent_\_unknown common_matcher" apply(simp add: packet_independent_\_unknown_def) apply(clarify) apply(rename_tac a p1 p2) apply(case_tac a) apply(simp_all add: bool_to_ternary_Unknown) apply(rename_tac l4ports, case_tac l4ports; simp add: bool_to_ternary_Unknown; fail)+ done lemma primitive_matcher_generic_common_matcher: "primitive_matcher_generic common_matcher" by unfold_locales simp_all (* What if we specify a port range where the start port is greater than the end port? For example, mathematically, {3 .. 2} = {}. Does iptables have the same behavior? For example, --source-port 1:0 raises an error on my system. For normal port specification, -m tcp, and -m multiport. There is also a manpage which states "if the first port is greater than the second one they will be swapped." I also saw a system where such an empty port range (--source-port 1:0) was really this impossible range and caused a rule that could never match. Because \ port. port \ {}. The behaviour if the end of the port range is smaller than the start is not 100% consistent among iptables versions and different modules. In general, it would be best to raise an error if such a range occurs. *) text\Warning: beware of the sloppy term `empty' portrange\ text\An `empty' port range means it can never match! Basically, @{term "MatchNot (Match (Src_Ports (L4Ports proto [(0,65535)])))"} is False\ lemma "\ matches (common_matcher, \) (MatchNot (Match (Src_Ports (L4Ports TCP [(0,65535)])))) a \p_iiface = ''eth0'', p_oiface = ''eth1'', p_src = ipv4addr_of_dotdecimal (192,168,2,45), p_dst= ipv4addr_of_dotdecimal (173,194,112,111), p_proto=TCP, p_sport=2065, p_dport=80, p_tcp_flags = {}, p_payload = '''', p_tag_ctstate = CT_New\" by(simp add: primitive_matcher_generic_common_matcher primitive_matcher_generic.Ports_single_not) text\An `empty' port range means it always matches! Basically, @{term "(MatchNot (Match (Src_Ports (L4Ports any []))))"} is True. This corresponds to firewall behavior, but usually you cannot specify an empty portrange in firewalls, but omission of portrange means no-port-restrictions, i.e. every port matches.\ lemma "matches (common_matcher, \) (MatchNot (Match (Src_Ports (L4Ports any [])))) a \p_iiface = ''eth0'', p_oiface = ''eth1'', p_src = ipv4addr_of_dotdecimal (192,168,2,45), p_dst= ipv4addr_of_dotdecimal (173,194,112,111), p_proto=TCP, p_sport=2065, p_dport=80, p_tcp_flags = {}, p_payload = '''', p_tag_ctstate = CT_New\" by(simp add: primitive_matcher_generic_common_matcher primitive_matcher_generic.Ports_single_not) text\If not a corner case, portrange matching is straight forward.\ lemma "matches (common_matcher, \) (Match (Src_Ports (L4Ports TCP [(1024,4096), (9999, 65535)]))) a \p_iiface = ''eth0'', p_oiface = ''eth1'', p_src = ipv4addr_of_dotdecimal (192,168,2,45), p_dst= ipv4addr_of_dotdecimal (173,194,112,111), p_proto=TCP, p_sport=2065, p_dport=80, p_tcp_flags = {}, p_payload = '''', p_tag_ctstate = CT_New\" "\ matches (common_matcher, \) (Match (Src_Ports (L4Ports TCP [(1024,4096), (9999, 65535)]))) a \p_iiface = ''eth0'', p_oiface = ''eth1'', p_src = ipv4addr_of_dotdecimal (192,168,2,45), p_dst= ipv4addr_of_dotdecimal (173,194,112,111), p_proto=TCP, p_sport=5000, p_dport=80, p_tcp_flags = {}, p_payload = '''', p_tag_ctstate = CT_New\" "\matches (common_matcher, \) (MatchNot (Match (Src_Ports (L4Ports TCP [(1024,4096), (9999, 65535)])))) a \p_iiface = ''eth0'', p_oiface = ''eth1'', p_src = ipv4addr_of_dotdecimal (192,168,2,45), p_dst= ipv4addr_of_dotdecimal (173,194,112,111), p_proto=TCP, p_sport=2065, p_dport=80, p_tcp_flags = {}, p_payload = '''', p_tag_ctstate = CT_New\" by(simp_all add: primitive_matcher_generic_common_matcher primitive_matcher_generic.Ports_single_not primitive_matcher_generic.Ports_single) text\Lemmas when matching on @{term Src} or @{term Dst}\ lemma common_matcher_SrcDst_defined: "common_matcher (Src m) p \ TernaryUnknown" "common_matcher (Dst m) p \ TernaryUnknown" "common_matcher (Src_Ports ps) p \ TernaryUnknown" "common_matcher (Dst_Ports ps) p \ TernaryUnknown" "common_matcher (MultiportPorts ps) p \ TernaryUnknown" apply(case_tac [!] m, case_tac [!] ps) apply(simp_all add: bool_to_ternary_Unknown) done lemma common_matcher_SrcDst_defined_simp: "common_matcher (Src x) p \ TernaryFalse \ common_matcher (Src x) p = TernaryTrue" "common_matcher (Dst x) p \ TernaryFalse \ common_matcher (Dst x) p = TernaryTrue" apply (metis eval_ternary_Not.cases common_matcher_SrcDst_defined(1) ternaryvalue.distinct(1)) apply (metis eval_ternary_Not.cases common_matcher_SrcDst_defined(2) ternaryvalue.distinct(1)) done (*The primitive_matcher_generic does not know anything about IP addresses*) lemma match_simplematcher_SrcDst: "matches (common_matcher, \) (Match (Src X)) a p \ p_src p \ ipt_iprange_to_set X" "matches (common_matcher, \) (Match (Dst X)) a p \ p_dst p \ ipt_iprange_to_set X" by(simp_all add: match_raw_ternary bool_to_ternary_simps split: ternaryvalue.split) lemma match_simplematcher_SrcDst_not: "matches (common_matcher, \) (MatchNot (Match (Src X))) a p \ p_src p \ ipt_iprange_to_set X" "matches (common_matcher, \) (MatchNot (Match (Dst X))) a p \ p_dst p \ ipt_iprange_to_set X" apply(simp_all add: matches_case_ternaryvalue_tuple split: ternaryvalue.split) apply(case_tac [!] X) apply(simp_all add: bool_to_ternary_simps) done lemma common_matcher_SrcDst_Inter: "(\m\set X. matches (common_matcher, \) (Match (Src m)) a p) \ p_src p \ (\x\set X. ipt_iprange_to_set x)" "(\m\set X. matches (common_matcher, \) (Match (Dst m)) a p) \ p_dst p \ (\x\set X. ipt_iprange_to_set x)" by(simp_all add: match_raw_ternary bool_to_ternary_simps split: ternaryvalue.split) subsection\Basic optimisations\ text\Perform very basic optimization. Remove matches to primitives which are essentially @{const MatchAny}\ fun optimize_primitive_univ :: "'i::len common_primitive match_expr \ 'i common_primitive match_expr" where "optimize_primitive_univ (Match (Src (IpAddrNetmask _ 0))) = MatchAny" | "optimize_primitive_univ (Match (Dst (IpAddrNetmask _ 0))) = MatchAny" | (*missing: the other IPs ...*) "optimize_primitive_univ (Match (IIface iface)) = (if iface = ifaceAny then MatchAny else (Match (IIface iface)))" | "optimize_primitive_univ (Match (OIface iface)) = (if iface = ifaceAny then MatchAny else (Match (OIface iface)))" | (*missing: L4Ports. But this introduces a new match, which causes problems. "optimize_primitive_univ (Match (Src_Ports (L4Ports proto [(s, e)]))) = (if s = 0 \ e = 0xFFFF then (Match (Prot (Proto proto))) else (Match (Src_Ports (L4Ports proto [(s, e)]))))" | "optimize_primitive_univ (Match (Dst_Ports (L4Ports proto [(s, e)]))) = (if s = 0 \ e = 0xFFFF then (Match (Prot (Proto proto))) else (Match (Dst_Ports (L4Ports proto [(s, e)]))))" |*) "optimize_primitive_univ (Match (Prot ProtoAny)) = MatchAny" | "optimize_primitive_univ (Match (L4_Flags (TCP_Flags m c))) = (if m = {} \ c = {} then MatchAny else (Match (L4_Flags (TCP_Flags m c))))" | "optimize_primitive_univ (Match (CT_State ctstate)) = (if ctstate_is_UNIV ctstate then MatchAny else Match (CT_State ctstate))" | "optimize_primitive_univ (Match m) = Match m" | (*"optimize_primitive_univ (MatchNot (MatchNot m)) = (optimize_primitive_univ m)" | --"needed to preserve normalized condition"*) "optimize_primitive_univ (MatchNot m) = (MatchNot (optimize_primitive_univ m))" | (*"optimize_primitive_univ (MatchAnd (Match (Extra e1)) (Match (Extra e2))) = optimize_primitive_univ (Match (Extra (e1@'' ''@e2)))" | -- "can be done but normalization does not work afterwards"*) "optimize_primitive_univ (MatchAnd m1 m2) = MatchAnd (optimize_primitive_univ m1) (optimize_primitive_univ m2)" | "optimize_primitive_univ MatchAny = MatchAny" lemma optimize_primitive_univ_unchanged_primitives: "optimize_primitive_univ (Match a) = (Match a) \ optimize_primitive_univ (Match a) = MatchAny" by (induction "(Match a)" rule: optimize_primitive_univ.induct) (auto split: if_split_asm) lemma optimize_primitive_univ_correct_matchexpr: fixes m::"'i::len common_primitive match_expr" shows "matches (common_matcher, \) m = matches (common_matcher, \) (optimize_primitive_univ m)" proof(simp add: fun_eq_iff, clarify, rename_tac a p) fix a and p :: "('i::len, 'a) tagged_packet_scheme" - have "65535 = (max_word::16 word)" by simp + have "65535 = (-1::16 word)" by simp then have port_range: "\s e port. s = 0 \ e = 0xFFFF \ (port::16 word) \ 0xFFFF" by (simp only:) simp have "ternary_ternary_eval (map_match_tac common_matcher p m) = ternary_ternary_eval (map_match_tac common_matcher p (optimize_primitive_univ m))" apply(induction m rule: optimize_primitive_univ.induct) by(simp_all add: port_range match_ifaceAny ipset_from_cidr_0 ctstate_is_UNIV) (*by(fastforce intro: arg_cong[where f=bool_to_ternary])+ if we add pots again*) thus "matches (common_matcher, \) m a p = matches (common_matcher, \) (optimize_primitive_univ m) a p" by(rule matches_iff_apply_f) qed corollary optimize_primitive_univ_correct: "approximating_bigstep_fun (common_matcher, \) p (optimize_matches optimize_primitive_univ rs) s = approximating_bigstep_fun (common_matcher, \) p rs s" using optimize_matches optimize_primitive_univ_correct_matchexpr by metis subsection\Abstracting over unknowns\ text\remove @{const Extra} (i.e. @{const TernaryUnknown}) match expressions\ fun upper_closure_matchexpr :: "action \ 'i::len common_primitive match_expr \ 'i common_primitive match_expr" where "upper_closure_matchexpr _ MatchAny = MatchAny" | "upper_closure_matchexpr Accept (Match (Extra _)) = MatchAny" | "upper_closure_matchexpr Reject (Match (Extra _)) = MatchNot MatchAny" | "upper_closure_matchexpr Drop (Match (Extra _)) = MatchNot MatchAny" | "upper_closure_matchexpr _ (Match m) = Match m" | "upper_closure_matchexpr Accept (MatchNot (Match (Extra _))) = MatchAny" | "upper_closure_matchexpr Drop (MatchNot (Match (Extra _))) = MatchNot MatchAny" | "upper_closure_matchexpr Reject (MatchNot (Match (Extra _))) = MatchNot MatchAny" | "upper_closure_matchexpr a (MatchNot (MatchNot m)) = upper_closure_matchexpr a m" | "upper_closure_matchexpr a (MatchNot (MatchAnd m1 m2)) = (let m1' = upper_closure_matchexpr a (MatchNot m1); m2' = upper_closure_matchexpr a (MatchNot m2) in (if m1' = MatchAny \ m2' = MatchAny then MatchAny else if m1' = MatchNot MatchAny then m2' else if m2' = MatchNot MatchAny then m1' else MatchNot (MatchAnd (MatchNot m1') (MatchNot m2'))) )" | "upper_closure_matchexpr _ (MatchNot m) = MatchNot m" | "upper_closure_matchexpr a (MatchAnd m1 m2) = MatchAnd (upper_closure_matchexpr a m1) (upper_closure_matchexpr a m2)" lemma upper_closure_matchexpr_generic: "a = Accept \ a = Drop \ remove_unknowns_generic (common_matcher, in_doubt_allow) a m = upper_closure_matchexpr a m" by(induction a m rule: upper_closure_matchexpr.induct) (simp_all add: remove_unknowns_generic_simps2 bool_to_ternary_Unknown common_matcher_SrcDst_defined) fun lower_closure_matchexpr :: "action \ 'i::len common_primitive match_expr \ 'i common_primitive match_expr" where "lower_closure_matchexpr _ MatchAny = MatchAny" | "lower_closure_matchexpr Accept (Match (Extra _)) = MatchNot MatchAny" | "lower_closure_matchexpr Reject (Match (Extra _)) = MatchAny" | "lower_closure_matchexpr Drop (Match (Extra _)) = MatchAny" | "lower_closure_matchexpr _ (Match m) = Match m" | "lower_closure_matchexpr Accept (MatchNot (Match (Extra _))) = MatchNot MatchAny" | "lower_closure_matchexpr Drop (MatchNot (Match (Extra _))) = MatchAny" | "lower_closure_matchexpr Reject (MatchNot (Match (Extra _))) = MatchAny" | "lower_closure_matchexpr a (MatchNot (MatchNot m)) = lower_closure_matchexpr a m" | "lower_closure_matchexpr a (MatchNot (MatchAnd m1 m2)) = (let m1' = lower_closure_matchexpr a (MatchNot m1); m2' = lower_closure_matchexpr a (MatchNot m2) in (if m1' = MatchAny \ m2' = MatchAny then MatchAny else if m1' = MatchNot MatchAny then m2' else if m2' = MatchNot MatchAny then m1' else MatchNot (MatchAnd (MatchNot m1') (MatchNot m2'))) )" | "lower_closure_matchexpr _ (MatchNot m) = MatchNot m" | "lower_closure_matchexpr a (MatchAnd m1 m2) = MatchAnd (lower_closure_matchexpr a m1) (lower_closure_matchexpr a m2)" lemma lower_closure_matchexpr_generic: "a = Accept \ a = Drop \ remove_unknowns_generic (common_matcher, in_doubt_deny) a m = lower_closure_matchexpr a m" by(induction a m rule: lower_closure_matchexpr.induct) (simp_all add: remove_unknowns_generic_simps2 bool_to_ternary_Unknown common_matcher_SrcDst_defined) end diff --git a/thys/Iptables_Semantics/Primitive_Matchers/Protocols_Normalize.thy b/thys/Iptables_Semantics/Primitive_Matchers/Protocols_Normalize.thy --- a/thys/Iptables_Semantics/Primitive_Matchers/Protocols_Normalize.thy +++ b/thys/Iptables_Semantics/Primitive_Matchers/Protocols_Normalize.thy @@ -1,506 +1,506 @@ theory Protocols_Normalize imports Common_Primitive_Lemmas "../Common/Word_Upto" begin section\Optimizing Protocols\ section\Optimizing protocols in match expressions\ fun compress_pos_protocols :: "protocol list \ protocol option" where "compress_pos_protocols [] = Some ProtoAny" | "compress_pos_protocols [p] = Some p" | "compress_pos_protocols (p1#p2#ps) = (case simple_proto_conjunct p1 p2 of None \ None | Some p \ compress_pos_protocols (p#ps))" lemma compress_pos_protocols_Some: "compress_pos_protocols ps = Some proto \ match_proto proto p_prot \ (\ p \ set ps. match_proto p p_prot)" proof(induction ps rule: compress_pos_protocols.induct) case (3 p1 p2 pps) thus ?case apply(cases "simple_proto_conjunct p1 p2") apply(simp; fail) using simple_proto_conjunct_Some by(simp) qed(simp)+ lemma compress_pos_protocols_None: "compress_pos_protocols ps = None \ \ (\ proto \ set ps. match_proto proto p_prot)" proof(induction ps rule: compress_pos_protocols.induct) case (3 i1 i2 iis) thus ?case apply(cases "simple_proto_conjunct i1 i2") apply(simp_all) using simple_proto_conjunct_None apply blast using simple_proto_conjunct_Some by blast qed(simp)+ (*the intuition behind the compress_protocols*) lemma "simple_proto_conjunct (Proto p1) (Proto p2) \ None \ \pkt. match_proto (Proto p1) pkt \ match_proto (Proto p2) pkt" apply(subgoal_tac "p1 = p2") apply(simp) apply(simp split: if_split_asm) done lemma "simple_proto_conjunct p1 (Proto p2) \ None \ \pkt. match_proto (Proto p2) pkt \ match_proto p1 pkt" apply(cases p1) apply(simp) apply(simp split: if_split_asm) done definition compress_protocols :: "protocol negation_type list \ (protocol list \ protocol list) option" where "compress_protocols ps \ case (compress_pos_protocols (getPos ps)) of None \ None - | Some proto \ if ProtoAny \ set (getNeg ps) \ (\p \ {0..max_word}. Proto p \ set (getNeg ps)) then + | Some proto \ if ProtoAny \ set (getNeg ps) \ (\p \ {0..- 1}. Proto p \ set (getNeg ps)) then None else if proto = ProtoAny then Some ([], getNeg ps) else if (\p \ set (getNeg ps). simple_proto_conjunct proto p \ None) then None else \ \\proto\ is a \primitive_protocol\ here. This is strict equality match, e.g.\ \ \protocol must be TCP. Thus, we can remove all negative matches!\ Some ([proto], [])" (* It is kind of messy to find a definition that checks whether a match is the exhaustive list and is executable *) - lemma all_proto_hlp2: "ProtoAny \ a \ (\p \ {0..max_word}. Proto p \ a) \ + lemma all_proto_hlp2: "ProtoAny \ a \ (\p \ {0..- 1}. Proto p \ a) \ ProtoAny \ a \ a = {p. p \ ProtoAny}" proof - - have all_proto_hlp: "ProtoAny \ a \ (\p \ {0..max_word}. Proto p \ a) \ a = {p. p \ ProtoAny}" + have all_proto_hlp: "ProtoAny \ a \ (\p \ {0..- 1}. Proto p \ a) \ a = {p. p \ ProtoAny}" by(auto intro: protocol.exhaust) thus ?thesis by blast qed - lemma set_word8_word_upto: "{0..(max_word :: 8 word)} = set (word_upto 0 255)" + lemma set_word8_word_upto: "{0..(- 1 :: 8 word)} = set (word_upto 0 255)" proof - - have \0xFF = (max_word :: 8 word)\ + have \0xFF = (- 1 :: 8 word)\ by simp then show ?thesis by (simp only:) (auto simp add: word_upto_set_eq) qed - lemma "(\p \ {0..max_word}. Proto p \ set (getNeg ps)) \ + lemma "(\p \ {0..- 1}. Proto p \ set (getNeg ps)) \ ((\p \ set (word_upto 0 255). Proto p \ set (getNeg ps)))" by(simp add: set_word8_word_upto) lemma compress_protocols_code[code]: "compress_protocols ps = (case (compress_pos_protocols (getPos ps)) of None \ None | Some proto \ if ProtoAny \ set (getNeg ps) \ (\p \ set (word_upto 0 255). Proto p \ set (getNeg ps)) then None else if proto = ProtoAny then Some ([], getNeg ps) else if (\p \ set (getNeg ps). simple_proto_conjunct proto p \ None) then None else Some ([proto], []) )" unfolding compress_protocols_def using set_word8_word_upto by presburger (*fully optimized, i.e. we cannot compress it better*) lemma "compress_protocols ps = Some (ps_pos, ps_neg) \ \ p. ((\m\set ps_pos. match_proto m p) \ (\m\set ps_neg. \ match_proto m p))" apply(simp add: compress_protocols_def all_proto_hlp2 split: option.split_asm if_split_asm) apply(subgoal_tac "\p. (Proto p) \ set ps_neg") apply(elim exE) apply(rename_tac x2 p) apply(rule_tac x=p in exI) apply(blast elim: match_proto.elims) apply(auto intro: protocol.exhaust) done definition compress_normalize_protocols_step :: "'i::len common_primitive match_expr \ 'i common_primitive match_expr option" where "compress_normalize_protocols_step m \ compress_normalize_primitive (is_Prot, prot_sel) Prot compress_protocols m" lemma (in primitive_matcher_generic) compress_normalize_protocols_step_Some: assumes "normalized_nnf_match m" and "compress_normalize_protocols_step m = Some m'" shows "matches (\, \) m' a p \ matches (\, \) m a p" proof(rule compress_normalize_primitive_Some[OF assms(1) wf_disc_sel_common_primitive(7), of compress_protocols]) show "compress_normalize_primitive (is_Prot, prot_sel) Prot compress_protocols m = Some m'" using assms(2) by(simp add: compress_normalize_protocols_step_def) next fix ps ps_pos ps_neg show "compress_protocols ps = Some (ps_pos, ps_neg) \ matches (\, \) (alist_and (NegPos_map Prot ((map Pos ps_pos)@(map Neg ps_neg)))) a p \ matches (\, \) (alist_and (NegPos_map Prot ps)) a p" apply(simp add: compress_protocols_def) apply(simp add: bunch_of_lemmata_about_matches alist_and_append NegPos_map_append) apply(simp add: nt_match_list_matches[symmetric] nt_match_list_simp) apply(simp add: NegPos_map_simps Prot_single Prot_single_not) apply(case_tac "compress_pos_protocols (getPos ps)") apply(simp_all) apply(drule_tac p_prot="p_proto p" in compress_pos_protocols_Some) apply(simp split:if_split_asm) using simple_proto_conjunct_None by auto qed lemma (in primitive_matcher_generic) compress_normalize_protocols_step_None: assumes "normalized_nnf_match m" and "compress_normalize_protocols_step m = None" shows "\ matches (\, \) m a p" proof(rule compress_normalize_primitive_None[OF assms(1) wf_disc_sel_common_primitive(7), of "compress_protocols"]) show "compress_normalize_primitive (is_Prot, prot_sel) Prot compress_protocols m = None" using assms(2) by(simp add: compress_normalize_protocols_step_def) next fix ps have if_option_Some: "((if P then None else Some x) = Some y) = (\P \ x = y)" for P and x::protocol and y by simp show "compress_protocols ps = None \ \ matches (\, \) (alist_and (NegPos_map Prot ps)) a p" apply(simp add: compress_protocols_def) apply(simp add: nt_match_list_matches[symmetric] nt_match_list_simp) apply(simp add: NegPos_map_simps Prot_single Prot_single_not) apply(cases "compress_pos_protocols (getPos ps)") apply(simp_all) apply(drule_tac p_prot="p_proto p" in compress_pos_protocols_None) apply(simp; fail) apply(drule_tac p_prot="p_proto p" in compress_pos_protocols_Some) apply(simp split:if_split_asm) apply fastforce apply(elim bexE exE) apply(simp) apply(elim simple_proto_conjunct.elims) apply(simp; fail) apply(simp; fail) using if_option_Some by metis qed lemma compress_normalize_protocols_step_nnf: "normalized_nnf_match m \ compress_normalize_protocols_step m = Some m' \ normalized_nnf_match m'" unfolding compress_normalize_protocols_step_def using compress_normalize_primitive_nnf[OF wf_disc_sel_common_primitive(7)] by blast (*not needed, I want it to introduce prot when I import from L4Ports!*) lemma compress_normalize_protocols_step_not_introduces_Prot: "\ has_disc is_Prot m \ normalized_nnf_match m \ compress_normalize_protocols_step m = Some m' \ \ has_disc is_Prot m'" apply(simp add: compress_normalize_protocols_step_def) apply(drule compress_normalize_primitive_not_introduces_C[where m=m and C'=Prot]) apply(simp_all add: wf_disc_sel_common_primitive(7)) apply(simp add: compress_protocols_def split: if_splits) done lemma compress_normalize_protocols_step_not_introduces_Prot_negated: assumes notdisc: "\ has_disc_negated is_Prot False m" and nm: "normalized_nnf_match m" and some: "compress_normalize_protocols_step m = Some m'" shows "\ has_disc_negated is_Prot False m'" apply(rule compress_normalize_primitive_not_introduces_C_negated[OF notdisc wf_disc_sel_common_primitive(7) nm]) using some apply(simp add: compress_normalize_protocols_step_def) by(simp add: compress_protocols_def split: option.split_asm if_split_asm) lemma compress_normalize_protocols_step_hasdisc: "\ has_disc disc m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_protocols_step m = Some m' \ normalized_nnf_match m' \ \ has_disc disc m'" unfolding compress_normalize_protocols_step_def using compress_normalize_primitive_hasdisc[OF _ wf_disc_sel_common_primitive(7)] by blast lemma compress_normalize_protocols_step_hasdisc_negated: "\ has_disc_negated disc neg m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_protocols_step m = Some m' \ normalized_nnf_match m' \ \ has_disc_negated disc neg m'" unfolding compress_normalize_protocols_step_def using compress_normalize_primitive_hasdisc_negated[OF _ wf_disc_sel_common_primitive(7)] by blast lemma compress_normalize_protocols_step_preserves_normalized_n_primitive: "normalized_n_primitive (disc, sel) P m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_protocols_step m = Some m' \ normalized_nnf_match m' \ normalized_n_primitive (disc, sel) P m'" unfolding compress_normalize_protocols_step_def using compress_normalize_primitve_preserves_normalized_n_primitive[OF _ wf_disc_sel_common_primitive(7)] by blast lemma "case compress_normalize_protocols_step (MatchAnd (MatchAnd (MatchAnd (Match ((Prot (Proto TCP)):: 32 common_primitive)) (MatchNot (Match (Prot (Proto UDP))))) (Match (IIface (Iface ''eth1'')))) (Match (Prot (Proto TCP)))) of Some ps \ opt_MatchAny_match_expr ps = MatchAnd (Match (Prot (Proto 6))) (Match (IIface (Iface ''eth1'')))" by eval value[code] "compress_normalize_protocols_step (MatchAny:: 32 common_primitive match_expr)" subsection\Importing the matches on @{typ primitive_protocol} from @{const L4Ports}\ (* add protocols from positive L4 ports into optimization. *) definition import_protocols_from_ports :: "'i::len common_primitive match_expr \ 'i common_primitive match_expr" where "import_protocols_from_ports m \ (case primitive_extractor (is_Src_Ports, src_ports_sel) m of (srcpts, rst1) \ case primitive_extractor (is_Dst_Ports, dst_ports_sel) rst1 of (dstpts, rst2) \ MatchAnd (MatchAnd (MatchAnd (andfold_MatchExp (map (Match \ (Prot \ (case_ipt_l4_ports (\proto x. Proto proto)))) (getPos srcpts))) (andfold_MatchExp (map (Match \ (Prot \ (case_ipt_l4_ports (\proto x. Proto proto)))) (getPos dstpts))) ) (alist_and' (NegPos_map Src_Ports srcpts @ NegPos_map Dst_Ports dstpts)) ) rst2 )" text\The @{const Proto} and @{const L4Ports} match make the following match impossible:\ lemma "compress_normalize_protocols_step (import_protocols_from_ports (MatchAnd (MatchAnd (Match (Prot (Proto TCP):: 32 common_primitive)) (Match (Src_Ports (L4Ports UDP [(22,22)])))) (Match (IIface (Iface ''eth1''))))) = None" by eval (*unfolding the whole primitive_extractor*) lemma import_protocols_from_ports_erule: "normalized_nnf_match m \ P m \ (\srcpts rst1 dstpts rst2. normalized_nnf_match m \ \ \\P m \\ erule consumes only first argument\ primitive_extractor (is_Src_Ports, src_ports_sel) m = (srcpts, rst1) \ primitive_extractor (is_Dst_Ports, dst_ports_sel) rst1 = (dstpts, rst2) \ normalized_nnf_match rst1 \ normalized_nnf_match rst2 \ P (MatchAnd (MatchAnd (MatchAnd (andfold_MatchExp (map (Match \ (Prot \ (case_ipt_l4_ports (\proto x. Proto proto)))) (getPos srcpts))) (andfold_MatchExp (map (Match \ (Prot \ (case_ipt_l4_ports (\proto x. Proto proto)))) (getPos dstpts)))) (alist_and' (NegPos_map Src_Ports srcpts @ NegPos_map Dst_Ports dstpts))) rst2)) \ P (import_protocols_from_ports m)" apply(simp add: import_protocols_from_ports_def) apply(case_tac "primitive_extractor (is_Src_Ports, src_ports_sel) m", rename_tac srcpts rst1) apply(simp) apply(case_tac "primitive_extractor (is_Dst_Ports, dst_ports_sel) rst1", rename_tac dstpts rst2) apply(simp) apply(frule(1) primitive_extractor_correct(2)[OF _ wf_disc_sel_common_primitive(1)]) apply(frule(1) primitive_extractor_correct(2)[OF _ wf_disc_sel_common_primitive(2)]) apply simp done lemma (in primitive_matcher_generic) import_protocols_from_ports: assumes normalized: "normalized_nnf_match m" shows "matches (\, \) (import_protocols_from_ports m) a p \ matches (\, \) m a p" proof- have add_protocol: "matches (\, \) (andfold_MatchExp (map (Match \ (Prot \ (case_ipt_l4_ports (\proto x. Proto proto)))) (getPos as))) a p \ matches (\, \) (alist_and (NegPos_map C as)) a p \ matches (\, \) (alist_and (NegPos_map C as)) a p" if C: "C = Src_Ports \ C = Dst_Ports" for C as proof(induction as) case Nil thus ?case by(simp) next case (Cons x xs) show ?case proof(cases x) case Neg with Cons.IH show ?thesis apply(simp add: bunch_of_lemmata_about_matches) by blast next case (Pos portmatch) with Cons.IH show ?thesis apply(cases portmatch) apply(simp add: andfold_MatchExp_matches bunch_of_lemmata_about_matches) using Ports_single_rewrite_Prot C by blast qed qed from normalized show ?thesis apply - apply(erule import_protocols_from_ports_erule) apply(simp; fail) apply(subst primitive_extractor_correct(1)[OF normalized wf_disc_sel_common_primitive(1), where \="(\,\)" and a=a and p=p, symmetric]) apply(simp; fail) apply(drule(1) primitive_extractor_correct(1)[OF _ wf_disc_sel_common_primitive(2), where \="(\,\)" and a=a and p=p]) apply(simp add: bunch_of_lemmata_about_matches matches_alist_and_alist_and' alist_and_append) using add_protocol by blast qed lemma import_protocols_from_ports_nnf: "normalized_nnf_match m \ normalized_nnf_match (import_protocols_from_ports m)" proof - have hlp: "\m\set (map (Match \ (Prot \ (case_ipt_l4_ports (\proto x. Proto proto)))) ls). normalized_nnf_match m" for ls apply(induction ls) apply(simp) apply(rename_tac l ls, case_tac l) by(simp) show "normalized_nnf_match m \ normalized_nnf_match (import_protocols_from_ports m)" apply(rule import_protocols_from_ports_erule) apply(simp_all) apply(simp add: normalized_nnf_match_alist_and') apply(safe) apply(rule andfold_MatchExp_normalized_nnf, simp add: hlp)+ done qed lemma import_protocols_from_ports_not_introduces_Prot_negated: "normalized_nnf_match m \ \ has_disc_negated is_Prot False m \ \ has_disc_negated is_Prot False (import_protocols_from_ports m)" apply(erule(1) import_protocols_from_ports_erule) apply(simp) apply(intro conjI) using andfold_MatchExp_not_disc_negated_mapMatch[ where C="Prot \ case_ipt_l4_ports (\proto x. Proto proto)", simplified] apply blast using andfold_MatchExp_not_disc_negated_mapMatch[ where C="Prot \ case_ipt_l4_ports (\proto x. Proto proto)", simplified] apply blast apply(simp add: has_disc_negated_alist_and') using not_has_disc_negated_NegPos_map[where disc=is_Prot and C=Src_Ports, simplified] not_has_disc_negated_NegPos_map[where disc=is_Prot and C=Dst_Ports, simplified] apply blast apply(drule(1) primitive_extractor_correct(6)[OF _ wf_disc_sel_common_primitive(1), where neg=False]) apply(drule(1) primitive_extractor_correct(6)[OF _ wf_disc_sel_common_primitive(2), where neg=False]) by blast lemma import_protocols_from_ports_hasdisc: "normalized_nnf_match m \ \ has_disc disc m \ (\a. \ disc (Prot a)) \ normalized_nnf_match (import_protocols_from_ports m) \ \ has_disc disc (import_protocols_from_ports m)" apply(intro conjI) using import_protocols_from_ports_nnf apply blast apply(erule(1) import_protocols_from_ports_erule) apply(simp) apply(intro conjI) using andfold_MatchExp_not_disc_mapMatch[ where C="Prot \ case_ipt_l4_ports (\proto x. Proto proto)", simplified] apply blast using andfold_MatchExp_not_disc_mapMatch[ where C="Prot \ case_ipt_l4_ports (\proto x. Proto proto)", simplified] apply blast subgoal for srcpts rst1 dstpts rst2 apply(frule(2) primitive_extractor_reassemble_not_has_disc[OF wf_disc_sel_common_primitive(1)]) apply(subgoal_tac "\ has_disc disc rst1") prefer 2 apply(drule(1) primitive_extractor_correct(4)[OF _ wf_disc_sel_common_primitive(1)]) apply blast apply(drule(2) primitive_extractor_reassemble_not_has_disc[OF wf_disc_sel_common_primitive(2)]) using has_disc_alist_and'_append by blast apply(drule(1) primitive_extractor_correct(4)[OF _ wf_disc_sel_common_primitive(1)]) apply(drule(1) primitive_extractor_correct(4)[OF _ wf_disc_sel_common_primitive(2)]) apply blast done lemma import_protocols_from_ports_hasdisc_negated: "\ has_disc_negated disc False m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ normalized_nnf_match (import_protocols_from_ports m) \ \ has_disc_negated disc False (import_protocols_from_ports m)" apply(intro conjI) using import_protocols_from_ports_nnf apply blast apply(erule(1) import_protocols_from_ports_erule) apply(simp) apply(intro conjI) using andfold_MatchExp_not_disc_negated_mapMatch[ where C="Prot \ case_ipt_l4_ports (\proto x. Proto proto)", simplified] apply blast using andfold_MatchExp_not_disc_negated_mapMatch[ where C="Prot \ case_ipt_l4_ports (\proto x. Proto proto)", simplified] apply blast subgoal for srcpts rst1 dstpts rst2 apply(frule(2) primitive_extractor_reassemble_not_has_disc_negated[OF wf_disc_sel_common_primitive(1)]) apply(subgoal_tac "\ has_disc_negated disc False rst1") prefer 2 apply(drule(1) primitive_extractor_correct(6)[OF _ wf_disc_sel_common_primitive(1)]) apply blast apply(drule(2) primitive_extractor_reassemble_not_has_disc_negated[OF wf_disc_sel_common_primitive(2)]) using has_disc_negated_alist_and'_append by blast apply(drule(1) primitive_extractor_correct(6)[OF _ wf_disc_sel_common_primitive(1)]) apply(drule(1) primitive_extractor_correct(6)[OF _ wf_disc_sel_common_primitive(2)]) apply blast done lemma import_protocols_from_ports_preserves_normalized_n_primitive: "normalized_n_primitive (disc, sel) f m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ normalized_nnf_match (import_protocols_from_ports m) \ normalized_n_primitive (disc, sel) f (import_protocols_from_ports m)" apply(intro conjI) using import_protocols_from_ports_nnf apply blast apply(erule(1) import_protocols_from_ports_erule) apply(simp) apply(intro conjI) subgoal for srcpts rst1 dstpts rst2 apply(rule andfold_MatchExp_normalized_n_primitive) using normalized_n_primitive_impossible_map by blast subgoal for srcpts rst1 dstpts rst2 apply(rule andfold_MatchExp_normalized_n_primitive) using normalized_n_primitive_impossible_map by blast subgoal for srcpts rst1 dstpts rst2 apply(frule(2) primitive_extractor_reassemble_normalized_n_primitive[OF wf_disc_sel_common_primitive(1)]) apply(subgoal_tac "normalized_n_primitive (disc, sel) f rst1") prefer 2 apply(drule(1) primitive_extractor_correct(5)[OF _ wf_disc_sel_common_primitive(1)]) apply blast apply(drule(2) primitive_extractor_reassemble_normalized_n_primitive[OF wf_disc_sel_common_primitive(2)]) using normalized_n_primitive_alist_and'_append by blast apply(drule(1) primitive_extractor_correct(5)[OF _ wf_disc_sel_common_primitive(1)]) apply(drule(1) primitive_extractor_correct(5)[OF _ wf_disc_sel_common_primitive(2)]) apply blast done subsection\Putting things together\ definition compress_normalize_protocols :: "'i::len common_primitive match_expr \ 'i common_primitive match_expr option" where "compress_normalize_protocols m \ compress_normalize_protocols_step (import_protocols_from_ports m)" lemma (in primitive_matcher_generic) compress_normalize_protocols_Some: assumes "normalized_nnf_match m" and "compress_normalize_protocols m = Some m'" shows "matches (\, \) m' a p \ matches (\, \) m a p" using assms apply(simp add: compress_normalize_protocols_def) by (metis import_protocols_from_ports import_protocols_from_ports_nnf compress_normalize_protocols_step_Some) lemma (in primitive_matcher_generic) compress_normalize_protocols_None: assumes "normalized_nnf_match m" and "compress_normalize_protocols m = None" shows "\ matches (\, \) m a p" using assms apply(simp add: compress_normalize_protocols_def) by (metis import_protocols_from_ports import_protocols_from_ports_nnf compress_normalize_protocols_step_None) lemma compress_normalize_protocols_nnf: "normalized_nnf_match m \ compress_normalize_protocols m = Some m' \ normalized_nnf_match m'" apply(simp add: compress_normalize_protocols_def) by (metis import_protocols_from_ports_nnf compress_normalize_protocols_step_nnf) lemma compress_normalize_protocols_not_introduces_Prot_negated: assumes notdisc: "\ has_disc_negated is_Prot False m" and nm: "normalized_nnf_match m" and some: "compress_normalize_protocols m = Some m'" shows "\ has_disc_negated is_Prot False m'" using assms apply(simp add: compress_normalize_protocols_def) using import_protocols_from_ports_nnf import_protocols_from_ports_not_introduces_Prot_negated compress_normalize_protocols_step_not_introduces_Prot_negated by auto lemma compress_normalize_protocols_hasdisc: "\ has_disc disc m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_protocols m = Some m' \ normalized_nnf_match m' \ \ has_disc disc m'" apply(simp add: compress_normalize_protocols_def) using import_protocols_from_ports_hasdisc compress_normalize_protocols_step_hasdisc by blast lemma compress_normalize_protocols_hasdisc_negated: "\ has_disc_negated disc False m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_protocols m = Some m' \ normalized_nnf_match m' \ \ has_disc_negated disc False m'" (*original lemma allowed arbitrary neg*) apply(simp add: compress_normalize_protocols_def) apply(frule(2) import_protocols_from_ports_hasdisc_negated) using compress_normalize_protocols_step_hasdisc_negated by blast lemma compress_normalize_protocols_preserves_normalized_n_primitive: "normalized_n_primitive (disc, sel) P m \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_protocols m = Some m' \ normalized_nnf_match m' \ normalized_n_primitive (disc, sel) P m'" apply(simp add: compress_normalize_protocols_def) using import_protocols_from_ports_preserves_normalized_n_primitive compress_normalize_protocols_step_preserves_normalized_n_primitive by blast lemma "case compress_normalize_protocols (MatchAnd (MatchAnd (MatchAnd (Match ((Prot (Proto TCP)):: 32 common_primitive)) (MatchNot (Match (Prot (Proto UDP))))) (Match (IIface (Iface ''eth1'')))) (Match (Prot (Proto TCP)))) of Some ps \ opt_MatchAny_match_expr ps = MatchAnd (Match (Prot (Proto 6))) (Match (IIface (Iface ''eth1'')))" by eval (*too many MatchAny!*) value[code] "compress_normalize_protocols (MatchAny:: 32 common_primitive match_expr)" end diff --git a/thys/Iptables_Semantics/Semantics_Embeddings.thy b/thys/Iptables_Semantics/Semantics_Embeddings.thy --- a/thys/Iptables_Semantics/Semantics_Embeddings.thy +++ b/thys/Iptables_Semantics/Semantics_Embeddings.thy @@ -1,347 +1,347 @@ theory Semantics_Embeddings imports "Simple_Firewall/SimpleFw_Compliance" Matching_Embeddings Semantics "Semantics_Ternary/Semantics_Ternary" begin section\Semantics Embedding\ subsection\Tactic @{const in_doubt_allow}\ lemma iptables_bigstep_undecided_to_undecided_in_doubt_allow_approx: assumes agree: "matcher_agree_on_exact_matches \ \" and good: "good_ruleset rs" and semantics: "\,\,p\ \rs, Undecided\ \ Undecided" shows "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Undecided \ (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" proof - from semantics good show ?thesis proof(induction rs Undecided Undecided rule: iptables_bigstep_induct) case Skip thus ?case by(auto intro: approximating_bigstep.skip) next case Log thus ?case by(auto intro: approximating_bigstep.empty approximating_bigstep.log approximating_bigstep.nomatch) next case (Nomatch m a) with not_exact_match_in_doubt_allow_approx_match[OF agree] have "a \ Log \ a \ Empty \ a = Accept \ Matching_Ternary.matches (\, in_doubt_allow) m a p \ \ Matching_Ternary.matches (\, in_doubt_allow) m a p" by(simp add: good_ruleset_alt) blast thus ?case by(cases a) (auto intro: approximating_bigstep.empty approximating_bigstep.log approximating_bigstep.accept approximating_bigstep.nomatch) next case (Seq rs rs1 rs2 t) from Seq have "good_ruleset rs1" and "good_ruleset rs2" by(simp_all add: good_ruleset_append) also from Seq iptables_bigstep_to_undecided have "t = Undecided" by simp ultimately show ?case using Seq by(fastforce intro: approximating_bigstep.decision Semantics_Ternary.seq') qed(simp_all add: good_ruleset_def) qed lemma FinalAllow_approximating_in_doubt_allow: assumes agree: "matcher_agree_on_exact_matches \ \" and good: "good_ruleset rs" and semantics: "\,\,p\ \rs, Undecided\ \ Decision FinalAllow" shows "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" proof - from semantics good show ?thesis proof(induction rs Undecided "Decision FinalAllow" rule: iptables_bigstep_induct) case Allow thus ?case by (auto intro: agree approximating_bigstep.accept in_doubt_allow_allows_Accept) next case (Seq rs rs1 rs2 t) from Seq have good1: "good_ruleset rs1" and good2: "good_ruleset rs2" by(simp_all add: good_ruleset_append) show ?case proof(cases t) case Decision with Seq good1 good2 show "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" by (auto intro: approximating_bigstep.decision approximating_bigstep.seq dest: Semantics.decisionD) next case Undecided with iptables_bigstep_undecided_to_undecided_in_doubt_allow_approx[OF agree good1] Seq have "(\, in_doubt_allow),p\ \rs1, Undecided\ \\<^sub>\ Undecided \ (\, in_doubt_allow),p\ \rs1, Undecided\ \\<^sub>\ Decision FinalAllow" by simp with Undecided Seq good1 good2 show "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" by (auto intro: approximating_bigstep.seq Semantics_Ternary.seq' approximating_bigstep.decision) qed next case Call_result thus ?case by(simp add: good_ruleset_alt) qed qed corollary FinalAllows_subseteq_in_doubt_allow: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow} \ {p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow}" using FinalAllow_approximating_in_doubt_allow by (metis (lifting, full_types) Collect_mono) (*referenced by name in paper*) corollary new_packets_to_simple_firewall_overapproximation: defines "preprocess rs \ upper_closure (optimize_matches abstract_for_simple_firewall (upper_closure (packet_assume_new rs)))" and "newpkt p \ match_tcp_flags ipt_tcp_syn (p_tcp_flags p) \ p_tag_ctstate p = CT_New" fixes p :: "('i::len, 'pkt_ext) tagged_packet_scheme" assumes "matcher_agree_on_exact_matches \ common_matcher" and "simple_ruleset rs" shows "{p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p} \ {p. simple_fw (to_simple_firewall (preprocess rs)) p = Decision FinalAllow \ newpkt p}" proof - from assms(3) have "{p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p} \ {p. (common_matcher, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(drule_tac rs=rs and \=\ in FinalAllows_subseteq_in_doubt_allow) using simple_imp_good_ruleset assms(4) apply blast by blast thus ?thesis unfolding newpkt_def preprocess_def using transform_simple_fw_upper(2)[OF assms(4)] by blast qed lemma approximating_bigstep_undecided_to_undecided_in_doubt_allow_approx: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Undecided \ \,\,p\ \rs, Undecided\ \ Undecided \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided Undecided rule: approximating_bigstep_induct) apply(simp_all) apply (metis iptables_bigstep.skip) apply (metis iptables_bigstep.empty iptables_bigstep.log iptables_bigstep.nomatch) apply(simp split: ternaryvalue.split_asm add: matches_case_ternaryvalue_tuple) apply (metis in_doubt_allow_allows_Accept iptables_bigstep.nomatch matches_casesE ternaryvalue.distinct(1) ternaryvalue.distinct(5)) apply(case_tac a) apply(simp_all) apply (metis iptables_bigstep.drop iptables_bigstep.nomatch) apply (metis iptables_bigstep.log iptables_bigstep.nomatch) apply (metis iptables_bigstep.nomatch iptables_bigstep.reject) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply (metis iptables_bigstep.empty iptables_bigstep.nomatch) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_append,clarify) by (metis approximating_bigstep_to_undecided iptables_bigstep.decision iptables_bigstep.seq) lemma FinalDeny_approximating_in_doubt_allow: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided "Decision FinalDeny" rule: approximating_bigstep_induct) apply(simp_all) apply (metis action.distinct(1) action.distinct(5) deny not_exact_match_in_doubt_allow_approx_match) apply(simp add: good_ruleset_append, clarify) apply(case_tac t) apply(simp) apply(drule(2) approximating_bigstep_undecided_to_undecided_in_doubt_allow_approx[where \=\]) apply(erule disjE) apply (metis iptables_bigstep.seq) apply (metis iptables_bigstep.decision iptables_bigstep.seq) by (metis Decision_approximating_bigstep_fun approximating_semantics_imp_fun iptables_bigstep.decision iptables_bigstep.seq) corollary FinalDenys_subseteq_in_doubt_allow: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ {p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalDeny}" using FinalDeny_approximating_in_doubt_allow by (metis (lifting, full_types) Collect_mono) text\ If our approximating firewall (the executable version) concludes that we deny a packet, the exact semantic agrees that this packet is definitely denied! \ corollary "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ approximating_bigstep_fun (\, in_doubt_allow) p rs Undecided = (Decision FinalDeny) \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny" apply(frule(1) FinalDeny_approximating_in_doubt_allow[where p=p and \=\]) apply(rule approximating_fun_imp_semantics) apply (metis good_imp_wf_ruleset) apply(simp_all) done subsection\Tactic @{const in_doubt_deny}\ lemma iptables_bigstep_undecided_to_undecided_in_doubt_deny_approx: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ \,\,p\ \rs, Undecided\ \ Undecided \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Undecided \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided Undecided rule: iptables_bigstep_induct) apply(simp_all) apply (metis approximating_bigstep.skip) apply (metis approximating_bigstep.empty approximating_bigstep.log approximating_bigstep.nomatch) apply(case_tac "a = Log") apply (metis approximating_bigstep.log approximating_bigstep.nomatch) apply(case_tac "a = Empty") apply (metis approximating_bigstep.empty approximating_bigstep.nomatch) apply(drule_tac a=a in not_exact_match_in_doubt_deny_approx_match) apply(simp_all) apply(simp add: good_ruleset_alt) apply fast apply (metis approximating_bigstep.drop approximating_bigstep.nomatch approximating_bigstep.reject) apply(frule iptables_bigstep_to_undecided) apply(simp) apply(simp add: good_ruleset_append) apply (metis (hide_lams, no_types) approximating_bigstep.decision Semantics_Ternary.seq') apply(simp add: good_ruleset_def) apply(simp add: good_ruleset_def) done lemma FinalDeny_approximating_in_doubt_deny: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided "Decision FinalDeny" rule: iptables_bigstep_induct) apply(simp_all) apply (metis approximating_bigstep.drop approximating_bigstep.reject in_doubt_deny_denies_DropReject) apply(case_tac t) apply(simp_all) prefer 2 apply(simp add: good_ruleset_append) apply (metis approximating_bigstep.decision approximating_bigstep.seq Semantics.decisionD state.inject) apply(simp add: good_ruleset_append, clarify) apply(drule(2) iptables_bigstep_undecided_to_undecided_in_doubt_deny_approx) apply(erule disjE) apply (metis approximating_bigstep.seq) apply (metis approximating_bigstep.decision Semantics_Ternary.seq') apply(simp add: good_ruleset_alt) done lemma approximating_bigstep_undecided_to_undecided_in_doubt_deny_approx: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Undecided \ \,\,p\ \rs, Undecided\ \ Undecided \ \,\,p\ \rs, Undecided\ \ Decision FinalAllow" apply(rotate_tac 2) apply(induction rs Undecided Undecided rule: approximating_bigstep_induct) apply(simp_all) apply (metis iptables_bigstep.skip) apply (metis iptables_bigstep.empty iptables_bigstep.log iptables_bigstep.nomatch) apply(simp split: ternaryvalue.split_asm add: matches_case_ternaryvalue_tuple) apply (metis in_doubt_allow_allows_Accept iptables_bigstep.nomatch matches_casesE ternaryvalue.distinct(1) ternaryvalue.distinct(5)) apply(case_tac a) apply(simp_all) apply (metis iptables_bigstep.accept iptables_bigstep.nomatch) apply (metis iptables_bigstep.log iptables_bigstep.nomatch) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply (metis iptables_bigstep.empty iptables_bigstep.nomatch) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_append,clarify) by (metis approximating_bigstep_to_undecided iptables_bigstep.decision iptables_bigstep.seq) lemma FinalAllow_approximating_in_doubt_deny: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ \,\,p\ \rs, Undecided\ \ Decision FinalAllow" apply(rotate_tac 2) apply(induction rs Undecided "Decision FinalAllow" rule: approximating_bigstep_induct) apply(simp_all) apply (metis action.distinct(1) action.distinct(5) iptables_bigstep.accept not_exact_match_in_doubt_deny_approx_match) apply(simp add: good_ruleset_append, clarify) apply(case_tac t) apply(simp) apply(drule(2) approximating_bigstep_undecided_to_undecided_in_doubt_deny_approx[where \=\]) apply(erule disjE) apply (metis iptables_bigstep.seq) apply (metis iptables_bigstep.decision iptables_bigstep.seq) by (metis Decision_approximating_bigstep_fun approximating_semantics_imp_fun iptables_bigstep.decision iptables_bigstep.seq) corollary FinalAllows_subseteq_in_doubt_deny: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ {p. (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow}" using FinalAllow_approximating_in_doubt_deny by (metis (lifting, full_types) Collect_mono) corollary new_packets_to_simple_firewall_underapproximation: defines "preprocess rs \ lower_closure (optimize_matches abstract_for_simple_firewall (lower_closure (packet_assume_new rs)))" and "newpkt p \ match_tcp_flags ipt_tcp_syn (p_tcp_flags p) \ p_tag_ctstate p = CT_New" fixes p :: "('i::len, 'pkt_ext) tagged_packet_scheme" assumes "matcher_agree_on_exact_matches \ common_matcher" and "simple_ruleset rs" shows "{p. simple_fw (to_simple_firewall (preprocess rs)) p = Decision FinalAllow \ newpkt p} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p}" proof - from assms(3) have "{p. (common_matcher, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p}" apply(drule_tac rs=rs and \=\ in FinalAllows_subseteq_in_doubt_deny) using simple_imp_good_ruleset assms(4) apply blast by blast thus ?thesis unfolding newpkt_def preprocess_def using transform_simple_fw_lower(2)[OF assms(4)] by blast qed subsection\Approximating Closures\ theorem FinalAllowClosure: assumes "matcher_agree_on_exact_matches \ \" and "good_ruleset rs" shows "{p. (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow}" and "{p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow} \ {p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow}" apply (metis FinalAllows_subseteq_in_doubt_deny assms) by (metis FinalAllows_subseteq_in_doubt_allow assms) theorem FinalDenyClosure: assumes "matcher_agree_on_exact_matches \ \" and "good_ruleset rs" shows "{p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalDeny}" and "{p. \,\,p\ \rs, Undecided\ \ Decision FinalDeny} \ {p. (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny}" apply (metis FinalDenys_subseteq_in_doubt_allow assms) by (metis FinalDeny_approximating_in_doubt_deny assms mem_Collect_eq subsetI) subsection\Exact Embedding\ lemma LukassLemma: assumes agree: "matcher_agree_on_exact_matches \ \" and noUnknown: "(\ r \ set rs. ternary_ternary_eval (map_match_tac \ p (get_match r)) \ TernaryUnknown)" and good: "good_ruleset rs" shows "(\,\),p\ \rs, s\ \\<^sub>\ t \ \,\,p\ \rs, s\ \ t" proof - { fix t \ \if we show it for arbitrary @{term t}, we can reuse this fact for the other direction.\ assume a: "(\,\),p\ \rs, s\ \\<^sub>\ t" from a good agree noUnknown have "\,\,p\ \rs, s\ \ t" proof(induction rs s t rule: approximating_bigstep_induct) qed(auto intro: approximating_bigstep.intros iptables_bigstep.intros dest: iptables_bigstepD dest: matches_comply_exact simp: good_ruleset_append) } note 1=this { assume a: "\,\,p\ \rs, s\ \ t" obtain x where "approximating_bigstep_fun (\,\) p rs s = x" by simp with approximating_fun_imp_semantics[OF good_imp_wf_ruleset[OF good]] have x: "(\,\),p\ \rs, s\ \\<^sub>\ x" by fast with 1 have "\,\,p\ \rs, s\ \ x" by simp with a iptables_bigstep_deterministic have "x = t" by metis hence "(\,\),p\ \rs, s\ \\<^sub>\ t" using x by blast } note 2=this from 1 2 show ?thesis by blast qed text\ For rulesets without @{term Call}s, the approximating ternary semantics can perfectly simulate the Boolean semantics. \ theorem \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_approximating_bigstep_iff_iptables_bigstep: assumes "\r \ set rs. \c. get_action r \ Call c" shows "((\\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c \),\),p\ \rs, s\ \\<^sub>\ t \ \,\,p\ \rs, s\ \ t" apply(rule iffI) apply(induction rs s t rule: approximating_bigstep_induct) apply(auto intro: iptables_bigstep.intros simp: \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_matching)[7] apply(insert assms) apply(induction rs s t rule: iptables_bigstep_induct) apply(auto intro: approximating_bigstep.intros simp: \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_matching) done corollary \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_approximating_bigstep_fun_iff_iptables_bigstep: assumes "good_ruleset rs" shows "approximating_bigstep_fun (\\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c \,\) p rs s = t \ \,\,p\ \rs, s\ \ t" apply(subst approximating_semantics_iff_fun_good_ruleset[symmetric]) using assms apply simp apply(subst \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_approximating_bigstep_iff_iptables_bigstep[where \=\]) using assms apply (simp add: good_ruleset_def) by simp text\The function @{const optimize_primitive_univ} was only applied to the ternary semantics. It is, in fact, also correct for the Boolean semantics, assuming the @{const common_matcher}.\ lemma Semantics_optimize_primitive_univ_common_matcher: assumes "matcher_agree_on_exact_matches \ common_matcher" shows "Semantics.matches \ (optimize_primitive_univ m) p = Semantics.matches \ m p" proof - - have "65535 = (max_word::16 word)" + have "65535 = (- 1::16 word)" by simp then have port_range: "\s e port. s = 0 \ e = 0xFFFF \ (port::16 word) \ 0xFFFF" by (simp only:) simp from assms show ?thesis apply(induction m rule: optimize_primitive_univ.induct) apply(auto elim!: matcher_agree_on_exact_matches_gammaE simp add: port_range match_ifaceAny ipset_from_cidr_0 ctstate_is_UNIV) done qed end diff --git a/thys/Iptables_Semantics/Simple_Firewall/SimpleFw_Compliance.thy b/thys/Iptables_Semantics/Simple_Firewall/SimpleFw_Compliance.thy --- a/thys/Iptables_Semantics/Simple_Firewall/SimpleFw_Compliance.thy +++ b/thys/Iptables_Semantics/Simple_Firewall/SimpleFw_Compliance.thy @@ -1,927 +1,927 @@ section\Iptables to Simple Firewall and Vice Versa\ theory SimpleFw_Compliance imports Simple_Firewall.SimpleFw_Semantics "../Primitive_Matchers/Transform" "../Primitive_Matchers/Primitive_Abstract" begin subsection\Simple Match to MatchExpr\ fun simple_match_to_ipportiface_match :: "'i::len simple_match \ 'i common_primitive match_expr" where "simple_match_to_ipportiface_match \iiface=iif, oiface=oif, src=sip, dst=dip, proto=p, sports=sps, dports=dps \ = MatchAnd (Match (IIface iif)) (MatchAnd (Match (OIface oif)) (MatchAnd (Match (Src (uncurry IpAddrNetmask sip))) (MatchAnd (Match (Dst (uncurry IpAddrNetmask dip))) (case p of ProtoAny \ MatchAny | Proto prim_p \ (MatchAnd (Match (Prot p)) (MatchAnd (Match (Src_Ports (L4Ports prim_p [sps]))) (Match (Dst_Ports (L4Ports prim_p [dps]))) )) ))))" lemma ports_to_set_singleton_simple_match_port: "p \ ports_to_set [a] \ simple_match_port a p" by(cases a, simp) theorem simple_match_to_ipportiface_match_correct: assumes valid: "simple_match_valid sm" shows "matches (common_matcher, \) (simple_match_to_ipportiface_match sm) a p \ simple_matches sm p" proof - obtain iif oif sip dip pro sps dps where sm: "sm = \iiface = iif, oiface = oif, src = sip, dst = dip, proto = pro, sports = sps, dports = dps\" by (cases sm) { fix ip have "p_src p \ ipt_iprange_to_set (uncurry IpAddrNetmask ip) \ simple_match_ip ip (p_src p)" and "p_dst p \ ipt_iprange_to_set (uncurry IpAddrNetmask ip) \ simple_match_ip ip (p_dst p)" by(simp split: uncurry_split)+ } note simple_match_ips=this { fix ps have "p_sport p \ ports_to_set [ps] \ simple_match_port ps (p_sport p)" and "p_dport p \ ports_to_set [ps] \ simple_match_port ps (p_dport p)" apply(case_tac [!] ps) by(simp_all) } note simple_match_ports=this from valid sm have valid':"pro = ProtoAny \ simple_match_port sps (p_sport p) \ simple_match_port dps (p_dport p)" apply(simp add: simple_match_valid_def) by blast show ?thesis unfolding sm apply(cases pro) subgoal apply(simp add: bunch_of_lemmata_about_matches simple_matches.simps) apply(simp add: match_raw_bool ternary_to_bool_bool_to_ternary simple_match_ips simple_match_ports simple_matches.simps) using valid' by simp apply(simp add: bunch_of_lemmata_about_matches simple_matches.simps) apply(simp add: match_raw_bool ternary_to_bool_bool_to_ternary simple_match_ips simple_match_ports simple_matches.simps) apply fast done qed subsection\MatchExpr to Simple Match\ fun common_primitive_match_to_simple_match :: "'i::len common_primitive match_expr \ 'i simple_match option" where "common_primitive_match_to_simple_match MatchAny = Some (simple_match_any)" | "common_primitive_match_to_simple_match (MatchNot MatchAny) = None" | "common_primitive_match_to_simple_match (Match (IIface iif)) = Some (simple_match_any\ iiface := iif \)" | "common_primitive_match_to_simple_match (Match (OIface oif)) = Some (simple_match_any\ oiface := oif \)" | "common_primitive_match_to_simple_match (Match (Src (IpAddrNetmask pre len))) = Some (simple_match_any\ src := (pre, len) \)" | "common_primitive_match_to_simple_match (Match (Dst (IpAddrNetmask pre len))) = Some (simple_match_any\ dst := (pre, len) \)" | "common_primitive_match_to_simple_match (Match (Prot p)) = Some (simple_match_any\ proto := p \)" | "common_primitive_match_to_simple_match (Match (Src_Ports (L4Ports p []))) = None" | "common_primitive_match_to_simple_match (Match (Src_Ports (L4Ports p [(s,e)]))) = Some (simple_match_any\ proto := Proto p, sports := (s,e) \)" | "common_primitive_match_to_simple_match (Match (Dst_Ports (L4Ports p []))) = None" | "common_primitive_match_to_simple_match (Match (Dst_Ports (L4Ports p [(s,e)]))) = Some (simple_match_any\ proto := Proto p, dports := (s,e) \)" | "common_primitive_match_to_simple_match (MatchNot (Match (Prot ProtoAny))) = None" | "common_primitive_match_to_simple_match (MatchAnd m1 m2) = (case (common_primitive_match_to_simple_match m1, common_primitive_match_to_simple_match m2) of (None, _) \ None | (_, None) \ None | (Some m1', Some m2') \ simple_match_and m1' m2')" | \ \undefined cases, normalize before!\ "common_primitive_match_to_simple_match (Match (Src (IpAddr _))) = undefined" | "common_primitive_match_to_simple_match (Match (Src (IpAddrRange _ _))) = undefined" | "common_primitive_match_to_simple_match (Match (Dst (IpAddr _))) = undefined" | "common_primitive_match_to_simple_match (Match (Dst (IpAddrRange _ _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (Prot _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (IIface _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (OIface _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (Src _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (Dst _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (MatchAnd _ _)) = undefined" | "common_primitive_match_to_simple_match (MatchNot (MatchNot _)) = undefined" | "common_primitive_match_to_simple_match (Match (Src_Ports _)) = undefined" | "common_primitive_match_to_simple_match (Match (Dst_Ports _)) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (Src_Ports _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (Dst_Ports _))) = undefined" | "common_primitive_match_to_simple_match (Match (CT_State _)) = undefined" | "common_primitive_match_to_simple_match (Match (L4_Flags _)) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (L4_Flags _))) = undefined" | "common_primitive_match_to_simple_match (Match (Extra _)) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (Extra _))) = undefined" | "common_primitive_match_to_simple_match (MatchNot (Match (CT_State _))) = undefined" subsubsection\Normalizing Interfaces\ text\As for now, negated interfaces are simply not allowed\ definition normalized_ifaces :: "'i::len common_primitive match_expr \ bool" where "normalized_ifaces m \ \ has_disc_negated (\a. is_Iiface a \ is_Oiface a) False m" subsubsection\Normalizing Protocols\ text\As for now, negated protocols are simply not allowed\ definition normalized_protocols :: "'i::len common_primitive match_expr \ bool" where "normalized_protocols m \ \ has_disc_negated is_Prot False m" lemma match_iface_simple_match_any_simps: "match_iface (iiface simple_match_any) (p_iiface p)" "match_iface (oiface simple_match_any) (p_oiface p)" "simple_match_ip (src simple_match_any) (p_src p)" "simple_match_ip (dst simple_match_any) (p_dst p)" "match_proto (proto simple_match_any) (p_proto p)" "simple_match_port (sports simple_match_any) (p_sport p)" "simple_match_port (dports simple_match_any) (p_dport p)" apply (simp_all add: simple_match_any_def match_ifaceAny ipset_from_cidr_0) - apply (subgoal_tac [!] "(65535::16 word) = max_word") + apply (subgoal_tac [!] "(65535::16 word) = - 1") apply (simp_all only:) apply simp_all done theorem common_primitive_match_to_simple_match: assumes "normalized_src_ports m" and "normalized_dst_ports m" and "normalized_src_ips m" and "normalized_dst_ips m" and "normalized_ifaces m" and "normalized_protocols m" and "\ has_disc is_L4_Flags m" and "\ has_disc is_CT_State m" and "\ has_disc is_MultiportPorts m" and "\ has_disc is_Extra m" shows "(Some sm = common_primitive_match_to_simple_match m \ matches (common_matcher, \) m a p \ simple_matches sm p) \ (common_primitive_match_to_simple_match m = None \ \ matches (common_matcher, \) m a p)" proof - show ?thesis using assms proof(induction m arbitrary: sm rule: common_primitive_match_to_simple_match.induct) case 1 thus ?case by(simp add: match_iface_simple_match_any_simps bunch_of_lemmata_about_matches simple_matches.simps) next case (9 p s e) thus ?case apply(simp add: match_iface_simple_match_any_simps simple_matches.simps) apply(simp add: match_raw_bool ternary_to_bool_bool_to_ternary) by fastforce next case 11 thus ?case apply(simp add: match_iface_simple_match_any_simps simple_matches.simps) apply(simp add: match_raw_bool ternary_to_bool_bool_to_ternary) by fastforce next case (13 m1 m2) let ?caseSome="Some sm = common_primitive_match_to_simple_match (MatchAnd m1 m2)" let ?caseNone="common_primitive_match_to_simple_match (MatchAnd m1 m2) = None" let ?goal="(?caseSome \ matches (common_matcher, \) (MatchAnd m1 m2) a p = simple_matches sm p) \ (?caseNone \ \ matches (common_matcher, \) (MatchAnd m1 m2) a p)" from 13 have normalized: "normalized_src_ports m1" "normalized_src_ports m2" "normalized_dst_ports m1" "normalized_dst_ports m2" "normalized_src_ips m1" "normalized_src_ips m2" "normalized_dst_ips m1" "normalized_dst_ips m2" "normalized_ifaces m1" "normalized_ifaces m2" "\ has_disc is_L4_Flags m1" "\ has_disc is_L4_Flags m2" "\ has_disc is_CT_State m1" "\ has_disc is_CT_State m2" "\ has_disc is_MultiportPorts m1" "\ has_disc is_MultiportPorts m2" "\ has_disc is_Extra m1" "\ has_disc is_Extra m2" "normalized_protocols m1" "normalized_protocols m2" by(simp_all add: normalized_protocols_def normalized_ifaces_def) { assume caseNone: ?caseNone { fix sm1 sm2 assume sm1: "common_primitive_match_to_simple_match m1 = Some sm1" and sm2: "common_primitive_match_to_simple_match m2 = Some sm2" and sma: "simple_match_and sm1 sm2 = None" from sma have 1: "\ (simple_matches sm1 p \ simple_matches sm2 p)" by (simp add: simple_match_and_correct) from normalized sm1 sm2 "13.IH" have 2: "(matches (common_matcher, \) m1 a p \ simple_matches sm1 p) \ (matches (common_matcher, \) m2 a p \ simple_matches sm2 p)" by force hence 2: "matches (common_matcher, \) (MatchAnd m1 m2) a p \ simple_matches sm1 p \ simple_matches sm2 p" by(simp add: bunch_of_lemmata_about_matches) from 1 2 have "\ matches (common_matcher, \) (MatchAnd m1 m2) a p" by blast } with caseNone have "common_primitive_match_to_simple_match m1 = None \ common_primitive_match_to_simple_match m2 = None \ \ matches (common_matcher, \) (MatchAnd m1 m2) a p" by(simp split:option.split_asm) hence "\ matches (common_matcher, \) (MatchAnd m1 m2) a p" apply(elim disjE) apply(simp_all) using "13.IH" normalized by(simp add: bunch_of_lemmata_about_matches)+ }note caseNone=this { assume caseSome: ?caseSome hence "\ sm1. common_primitive_match_to_simple_match m1 = Some sm1" and "\ sm2. common_primitive_match_to_simple_match m2 = Some sm2" by(simp_all split: option.split_asm) from this obtain sm1 sm2 where sm1: "Some sm1 = common_primitive_match_to_simple_match m1" and sm2: "Some sm2 = common_primitive_match_to_simple_match m2" by fastforce+ with "13.IH" normalized have "matches (common_matcher, \) m1 a p = simple_matches sm1 p \ matches (common_matcher, \) m2 a p = simple_matches sm2 p" by simp hence 1: "matches (common_matcher, \) (MatchAnd m1 m2) a p \ simple_matches sm1 p \ simple_matches sm2 p" by(simp add: bunch_of_lemmata_about_matches) from caseSome sm1 sm2 have "simple_match_and sm1 sm2 = Some sm" by(simp split: option.split_asm) hence 2: "simple_matches sm p \ simple_matches sm1 p \ simple_matches sm2 p" by(simp add: simple_match_and_correct) from 1 2 have "matches (common_matcher, \) (MatchAnd m1 m2) a p = simple_matches sm p" by simp } note caseSome=this from caseNone caseSome show ?goal by blast qed(simp_all add: match_iface_simple_match_any_simps simple_matches.simps normalized_protocols_def normalized_ifaces_def, simp_all add: bunch_of_lemmata_about_matches, simp_all add: match_raw_bool ternary_to_bool_bool_to_ternary) qed lemma simple_fw_remdups_Rev: "simple_fw (remdups_rev rs) p = simple_fw rs p" apply(induction rs p rule: simple_fw.induct) apply(simp add: remdups_rev_def) apply(simp_all add: remdups_rev_fst remdups_rev_removeAll simple_fw_not_matches_removeAll) done fun action_to_simple_action :: "action \ simple_action" where "action_to_simple_action action.Accept = simple_action.Accept" | "action_to_simple_action action.Drop = simple_action.Drop" | "action_to_simple_action _ = undefined" definition check_simple_fw_preconditions :: "'i::len common_primitive rule list \ bool" where "check_simple_fw_preconditions rs \ \r \ set rs. (case r of (Rule m a) \ normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ normalized_ifaces m \ normalized_protocols m \ \ has_disc is_L4_Flags m \ \ has_disc is_CT_State m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m \ (a = action.Accept \ a = action.Drop))" (*apart from MatchNot MatchAny, the normalizations imply nnf*) lemma "normalized_src_ports m \ normalized_nnf_match m" apply(induction m rule: normalized_src_ports.induct) apply(simp_all)[15] oops lemma "\ matcheq_matchNone m \ normalized_src_ports m \ normalized_nnf_match m" by(induction m rule: normalized_src_ports.induct) (simp_all) value "check_simple_fw_preconditions [Rule (MatchNot (MatchNot (MatchNot (Match (Src a))))) action.Accept]" definition to_simple_firewall :: "'i::len common_primitive rule list \ 'i simple_rule list" where "to_simple_firewall rs \ if check_simple_fw_preconditions rs then List.map_filter (\r. case r of Rule m a \ (case (common_primitive_match_to_simple_match m) of None \ None | Some sm \ Some (SimpleRule sm (action_to_simple_action a)))) rs else undefined" lemma to_simple_firewall_simps: "to_simple_firewall [] = []" "check_simple_fw_preconditions ((Rule m a)#rs) \ to_simple_firewall ((Rule m a)#rs) = (case common_primitive_match_to_simple_match m of None \ to_simple_firewall rs | Some sm \ (SimpleRule sm (action_to_simple_action a)) # to_simple_firewall rs)" "\ check_simple_fw_preconditions rs' \ to_simple_firewall rs' = undefined" by(auto simp add: to_simple_firewall_def List.map_filter_simps check_simple_fw_preconditions_def split: option.split) lemma "check_simple_fw_preconditions [Rule (MatchAnd (Match (Src (IpAddrNetmask (ipv4addr_of_dotdecimal (127, 0, 0, 0)) 8))) (MatchAnd (Match (Dst_Ports (L4Ports TCP [(0, 65535)]))) (Match (Src_Ports (L4Ports TCP [(0, 65535)]))))) Drop]" by eval lemma "to_simple_firewall [Rule (MatchAnd (Match (Src (IpAddrNetmask (ipv4addr_of_dotdecimal (127, 0, 0, 0)) 8))) (MatchAnd (Match (Dst_Ports (L4Ports TCP [(0, 65535)]))) (Match (Src_Ports (L4Ports TCP [(0, 65535)]))))) Drop] = [SimpleRule \iiface = Iface ''+'', oiface = Iface ''+'', src = (0x7F000000, 8), dst = (0, 0), proto = Proto 6, sports = (0, 0xFFFF), dports = (0, 0xFFFF)\ simple_action.Drop]" by eval lemma "check_simple_fw_preconditions [Rule (MatchAnd MatchAny MatchAny) Drop]" by(simp add: check_simple_fw_preconditions_def normalized_ifaces_def normalized_protocols_def) lemma "to_simple_firewall [Rule (MatchAnd MatchAny (MatchAny::32 common_primitive match_expr)) Drop] = [SimpleRule \iiface = Iface ''+'', oiface = Iface ''+'', src = (0, 0), dst = (0, 0), proto = ProtoAny, sports = (0, 0xFFFF), dports = (0, 0xFFFF)\ simple_action.Drop]" by eval lemma "to_simple_firewall [Rule (Match (Src (IpAddrNetmask (ipv4addr_of_dotdecimal (127, 0, 0, 0)) 8))) Drop] = [SimpleRule \iiface = Iface ''+'', oiface = Iface ''+'', src = (0x7F000000, 8), dst = (0, 0), proto = ProtoAny, sports = (0, 0xFFFF), dports = (0, 0xFFFF)\ simple_action.Drop]" by eval theorem to_simple_firewall: "check_simple_fw_preconditions rs \ approximating_bigstep_fun (common_matcher, \) p rs Undecided = simple_fw (to_simple_firewall rs) p" proof(induction rs) case Nil thus ?case by(simp add: to_simple_firewall_simps) next case (Cons r rs) from Cons have IH: "approximating_bigstep_fun (common_matcher, \) p rs Undecided = simple_fw (to_simple_firewall rs) p" by(simp add: check_simple_fw_preconditions_def) obtain m a where r: "r = Rule m a" by(cases r, simp) from Cons.prems have "check_simple_fw_preconditions [r]" by(simp add: check_simple_fw_preconditions_def) with r common_primitive_match_to_simple_match[where p = p] have match: "\ sm. common_primitive_match_to_simple_match m = Some sm \ matches (common_matcher, \) m a p = simple_matches sm p" and nomatch: "common_primitive_match_to_simple_match m = None \ \ matches (common_matcher, \) m a p" unfolding check_simple_fw_preconditions_def by simp_all from to_simple_firewall_simps r Cons.prems have to_simple_firewall_simps': "to_simple_firewall (Rule m a # rs) = (case common_primitive_match_to_simple_match m of None \ to_simple_firewall rs | Some sm \ SimpleRule sm (action_to_simple_action a) # to_simple_firewall rs)" by simp from \check_simple_fw_preconditions [r]\ have "a = action.Accept \ a = action.Drop" by(simp add: r check_simple_fw_preconditions_def) thus ?case by(auto simp add: r to_simple_firewall_simps' IH match nomatch split: option.split action.split) qed lemma ctstate_assume_new_not_has_CT_State: "r \ set (ctstate_assume_new rs) \ \ has_disc is_CT_State (get_match r)" apply(simp add: ctstate_assume_new_def) apply(induction rs) apply(simp add: optimize_matches_def; fail) apply(simp add: optimize_matches_def) apply(rename_tac r' rs, case_tac r') apply(safe) apply(simp add: split:if_split_asm) apply(elim disjE) apply(simp_all add: not_hasdisc_ctstate_assume_state split:if_split_asm) done text\The precondition for the simple firewall can be easily fulfilled. The subset relation is due to abstracting over some primitives (e.g., negated primitives, l4 flags)\ theorem transform_simple_fw_upper: defines "preprocess rs \ upper_closure (optimize_matches abstract_for_simple_firewall (upper_closure (packet_assume_new rs)))" and "newpkt p \ match_tcp_flags ipt_tcp_syn (p_tcp_flags p) \ p_tag_ctstate p = CT_New" assumes simplers: "simple_ruleset (rs:: 'i::len common_primitive rule list)" \ \the preconditions for the simple firewall are fulfilled, definitely no runtime failure\ shows "check_simple_fw_preconditions (preprocess rs)" \ \the set of new packets, which are accepted is an overapproximations\ and "{p. (common_matcher, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. simple_fw (to_simple_firewall (preprocess rs)) p = Decision FinalAllow \ newpkt p}" \ \Fun fact: The theorem holds for a tagged packet. The simple firewall just ignores the tag. You may explicitly untag, if you wish to, but a @{typ "'i tagged_packet"} is just an extension of the @{typ "'i simple_packet"} used by the simple firewall\ unfolding check_simple_fw_preconditions_def preprocess_def apply(clarify, rename_tac r, case_tac r, rename_tac m a, simp) proof - let ?rs2="upper_closure (packet_assume_new rs)" let ?rs3="optimize_matches abstract_for_simple_firewall ?rs2" let ?rs'="upper_closure ?rs3" let ?\="(common_matcher, in_doubt_allow) :: ('i::len common_primitive, ('i, 'a) tagged_packet_scheme) match_tac" let ?fw="\rs p. approximating_bigstep_fun ?\ p rs Undecided" from packet_assume_new_simple_ruleset[OF simplers] have s1: "simple_ruleset (packet_assume_new rs)" . from transform_upper_closure(2)[OF s1] have s2: "simple_ruleset ?rs2" . from s2 have s3: "simple_ruleset ?rs3" by (simp add: optimize_matches_simple_ruleset) from transform_upper_closure(2)[OF s3] have s4: "simple_ruleset ?rs'" . from transform_upper_closure(3)[OF s1] have nnf2: "\r\set (upper_closure (packet_assume_new rs)). normalized_nnf_match (get_match r)" by simp { fix m a assume r: "Rule m a \ set ?rs'" from s4 r have a: "(a = action.Accept \ a = action.Drop)" by(auto simp add: simple_ruleset_def) have "r \ set (packet_assume_new rs) \ \ has_disc is_CT_State (get_match r)" for r by(simp add: packet_assume_new_def ctstate_assume_new_not_has_CT_State) with transform_upper_closure(4)[OF s1, where disc=is_CT_State] have "\r\set (upper_closure (packet_assume_new rs)). \ has_disc is_CT_State (get_match r)" by simp with abstract_primitive_preserves_nodisc[where disc'="is_CT_State"] have "\r\set ?rs3. \ has_disc is_CT_State (get_match r)" apply(intro optimize_matches_preserves) by(auto simp add: abstract_for_simple_firewall_def) with transform_upper_closure(4)[OF s3, where disc=is_CT_State] have "\r\set ?rs'. \ has_disc is_CT_State (get_match r)" by simp with r have no_CT: "\ has_disc is_CT_State m" by fastforce from abstract_for_simple_firewall_hasdisc have "\r\set ?rs3. \ has_disc is_L4_Flags (get_match r)" by(intro optimize_matches_preserves, auto) with transform_upper_closure(4)[OF s3, where disc=is_L4_Flags] have "\r\set ?rs'. \ has_disc is_L4_Flags (get_match r)" by simp with r have no_L4_Flags: "\ has_disc is_L4_Flags m" by fastforce from nnf2 abstract_for_simple_firewall_negated_ifaces_prots have ifaces: "\r\set ?rs3. \ has_disc_negated (\a. is_Iiface a \ is_Oiface a) False (get_match r)" and protocols_rs3: "\r\set ?rs3. \ has_disc_negated is_Prot False (get_match r)" by(intro optimize_matches_preserves, blast)+ from ifaces have iface_in: "\r\set ?rs3. \ has_disc_negated is_Iiface False (get_match r)" and iface_out: "\r\set ?rs3. \ has_disc_negated is_Oiface False (get_match r)" using has_disc_negated_disj_split by blast+ from transform_upper_closure(3)[OF s3] have "\r\set ?rs'. normalized_nnf_match (get_match r) \ normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r) \ \ has_disc is_Extra (get_match r)" . with r have normalized: "normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ \ has_disc is_MultiportPorts m & \ has_disc is_Extra m" by fastforce (*things are complicated because upper closure could introduce negated protocols. should not happen if we don't have negated ports in it *) from transform_upper_closure(5)[OF s3] iface_in iface_out have "\r\set ?rs'. \ has_disc_negated is_Iiface False (get_match r) \ \ has_disc_negated is_Oiface False (get_match r)" by simp (*500ms*) with r have abstracted_ifaces: "normalized_ifaces m" unfolding normalized_ifaces_def has_disc_negated_disj_split by fastforce from transform_upper_closure(3)[OF s1] normalized_n_primitive_imp_not_disc_negated[OF wf_disc_sel_common_primitive(1)] normalized_n_primitive_imp_not_disc_negated[OF wf_disc_sel_common_primitive(2)] have "\r\ set ?rs2. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" apply(simp add: normalized_src_ports_def2 normalized_dst_ports_def2) by blast from this have "\r\set ?rs3. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" apply - apply(rule optimize_matches_preserves) apply(intro conjI) apply(intro abstract_for_simple_firewall_preserves_nodisc_negated, simp_all)+ by (simp add: abstract_for_simple_firewall_def abstract_primitive_preserves_nodisc) from this protocols_rs3 transform_upper_closure(5)[OF s3, where disc=is_Prot, simplified] have "\r\set ?rs'. \ has_disc_negated is_Prot False (get_match r)" by simp with r have abstracted_prots: "normalized_protocols m" unfolding normalized_protocols_def has_disc_negated_disj_split by fastforce from no_CT no_L4_Flags s4 normalized a abstracted_ifaces abstracted_prots show "normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ normalized_ifaces m \ normalized_protocols m \ \ has_disc is_L4_Flags m \ \ has_disc is_CT_State m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m \ (a = action.Accept \ a = action.Drop)" by(simp) } hence simple_fw_preconditions: "check_simple_fw_preconditions ?rs'" unfolding check_simple_fw_preconditions_def by(clarify, rename_tac r, case_tac r, rename_tac m a, simp) have 1: "{p. ?\,p\ \?rs', Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p. ?\,p\ \?rs3, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst transform_upper_closure(1)[OF s3]) by simp from abstract_primitive_in_doubt_allow_generic(2)[OF primitive_matcher_generic_common_matcher nnf2 s2] have 2: "{p. ?\,p\ \upper_closure (packet_assume_new rs), Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \?rs3, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" by(auto simp add: abstract_for_simple_firewall_def) have 3: "{p. ?\,p\ \upper_closure (packet_assume_new rs), Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst transform_upper_closure(1)[OF s1]) apply(subst approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF s1]]) apply(subst approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]]) using packet_assume_new newpkt_def by fastforce have 4: "\p. ?\,p\ \?rs', Undecided\ \\<^sub>\ Decision FinalAllow \ ?fw ?rs' p = Decision FinalAllow" using approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF s4]] by fast have "{p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \?rs', Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst 1) apply(subst 3[symmetric]) using 2 by blast thus "{p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. simple_fw (to_simple_firewall ?rs') p = Decision FinalAllow \ newpkt p}" apply safe subgoal for p using to_simple_firewall[OF simple_fw_preconditions, where p = p] 4 by auto done qed (*Copy&paste from transform_simple_fw_upper*) theorem transform_simple_fw_lower: defines "preprocess rs \ lower_closure (optimize_matches abstract_for_simple_firewall (lower_closure (packet_assume_new rs)))" and "newpkt p \ match_tcp_flags ipt_tcp_syn (p_tcp_flags p) \ p_tag_ctstate p = CT_New" assumes simplers: "simple_ruleset (rs:: 'i::len common_primitive rule list)" \ \the preconditions for the simple firewall are fulfilled, definitely no runtime failure\ shows "check_simple_fw_preconditions (preprocess rs)" \ \the set of new packets, which are accepted is an underapproximation\ and "{p. simple_fw (to_simple_firewall (preprocess rs)) p = Decision FinalAllow \ newpkt p} \ {p. (common_matcher, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" unfolding check_simple_fw_preconditions_def preprocess_def apply(clarify, rename_tac r, case_tac r, rename_tac m a, simp) proof - let ?rs2="lower_closure (packet_assume_new rs)" let ?rs3="optimize_matches abstract_for_simple_firewall ?rs2" let ?rs'="lower_closure ?rs3" let ?\="(common_matcher, in_doubt_deny) :: ('i::len common_primitive, ('i, 'a) tagged_packet_scheme) match_tac" let ?fw="\rs p. approximating_bigstep_fun ?\ p rs Undecided" from packet_assume_new_simple_ruleset[OF simplers] have s1: "simple_ruleset (packet_assume_new rs)" . from transform_lower_closure(2)[OF s1] have s2: "simple_ruleset (lower_closure (packet_assume_new rs))" . from s2 have s3: "simple_ruleset ?rs3" by (simp add: optimize_matches_simple_ruleset) from transform_lower_closure(2)[OF s3] have s4: "simple_ruleset ?rs'" . from transform_lower_closure(3)[OF s1] have nnf2: "\r\set (lower_closure (packet_assume_new rs)). normalized_nnf_match (get_match r)" by simp { fix m a assume r: "Rule m a \ set ?rs'" from s4 r have a: "(a = action.Accept \ a = action.Drop)" by(auto simp add: simple_ruleset_def) have "r \ set (packet_assume_new rs) \ \ has_disc is_CT_State (get_match r)" for r by(simp add: packet_assume_new_def ctstate_assume_new_not_has_CT_State) with transform_lower_closure(4)[OF s1, where disc=is_CT_State] have "\r\set (lower_closure (packet_assume_new rs)). \ has_disc is_CT_State (get_match r)" by fastforce with abstract_primitive_preserves_nodisc[where disc'="is_CT_State"] have "\r\set ?rs3. \ has_disc is_CT_State (get_match r)" apply(intro optimize_matches_preserves) by(auto simp add: abstract_for_simple_firewall_def) with transform_lower_closure(4)[OF s3, where disc=is_CT_State] have "\r\set ?rs'. \ has_disc is_CT_State (get_match r)" by fastforce with r have no_CT: "\ has_disc is_CT_State m" by fastforce from abstract_for_simple_firewall_hasdisc have "\r\set ?rs3. \ has_disc is_L4_Flags (get_match r)" by(intro optimize_matches_preserves, blast) with transform_lower_closure(4)[OF s3, where disc=is_L4_Flags] have "\r\set ?rs'. \ has_disc is_L4_Flags (get_match r)" by fastforce with r have no_L4_Flags: "\ has_disc is_L4_Flags m" by fastforce from nnf2 abstract_for_simple_firewall_negated_ifaces_prots have ifaces: "\r\set ?rs3. \ has_disc_negated (\a. is_Iiface a \ is_Oiface a) False (get_match r)" and protocols_rs3: "\r\set ?rs3. \ has_disc_negated is_Prot False (get_match r)" by(intro optimize_matches_preserves, blast)+ from ifaces have iface_in: "\r\set ?rs3. \ has_disc_negated is_Iiface False (get_match r)" and iface_out: "\r\set ?rs3. \ has_disc_negated is_Oiface False (get_match r)" using has_disc_negated_disj_split by blast+ from transform_lower_closure(3)[OF s3] have "\r\set ?rs'. normalized_nnf_match (get_match r) \ normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r) \ \ has_disc is_Extra (get_match r)" . with r have normalized: "normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m" by fastforce from transform_lower_closure(5)[OF s3] iface_in iface_out have "\r\set ?rs'. \ has_disc_negated is_Iiface False (get_match r) \ \ has_disc_negated is_Oiface False (get_match r)" by simp (*500ms*) with r have abstracted_ifaces: "normalized_ifaces m" unfolding normalized_ifaces_def has_disc_negated_disj_split by fastforce from transform_lower_closure(3)[OF s1] normalized_n_primitive_imp_not_disc_negated[OF wf_disc_sel_common_primitive(1)] normalized_n_primitive_imp_not_disc_negated[OF wf_disc_sel_common_primitive(2)] have "\r\set ?rs2. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" apply(simp add: normalized_src_ports_def2 normalized_dst_ports_def2) by blast from this have "\r\set ?rs3. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" apply - apply(rule optimize_matches_preserves) apply(intro conjI) apply(intro abstract_for_simple_firewall_preserves_nodisc_negated, simp_all)+ by (simp add: abstract_for_simple_firewall_def abstract_primitive_preserves_nodisc) from this protocols_rs3 transform_lower_closure(5)[OF s3, where disc=is_Prot, simplified] have "\r\set ?rs'. \ has_disc_negated is_Prot False (get_match r)" by simp with r have abstracted_prots: "normalized_protocols m" unfolding normalized_protocols_def has_disc_negated_disj_split by fastforce from no_CT no_L4_Flags s4 normalized a abstracted_ifaces abstracted_prots show "normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ normalized_ifaces m \ normalized_protocols m \ \ has_disc is_L4_Flags m \ \ has_disc is_CT_State m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m \ (a = action.Accept \ a = action.Drop)" by(simp) } hence simple_fw_preconditions: "check_simple_fw_preconditions ?rs'" unfolding check_simple_fw_preconditions_def by(clarify, rename_tac r, case_tac r, rename_tac m a, simp) have 1: "{p. ?\,p\ \?rs', Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p. ?\,p\ \?rs3, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst transform_lower_closure(1)[OF s3]) by simp from abstract_primitive_in_doubt_deny_generic(1)[OF primitive_matcher_generic_common_matcher nnf2 s2] have 2: "{p. ?\,p\ \?rs3, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \lower_closure (packet_assume_new rs), Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" by(auto simp add: abstract_for_simple_firewall_def) have 3: "{p. ?\,p\ \lower_closure (packet_assume_new rs), Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst transform_lower_closure(1)[OF s1]) apply(subst approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF s1]]) apply(subst approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]]) using packet_assume_new newpkt_def by fastforce have 4: "\p. ?\,p\ \?rs', Undecided\ \\<^sub>\ Decision FinalAllow \ ?fw ?rs' p = Decision FinalAllow" using approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF s4]] by fast have "{p. ?\,p\ \?rs', Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst 1) apply(subst 3[symmetric]) using 2 by blast thus "{p. simple_fw (to_simple_firewall ?rs') p = Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} " apply safe subgoal for p using to_simple_firewall[OF simple_fw_preconditions, where p = p] 4 by auto done qed definition "to_simple_firewall_without_interfaces ipassmt rtblo rs \ to_simple_firewall (upper_closure (optimize_matches (abstract_primitive (\r. case r of Pos a \ is_Iiface a \ is_Oiface a | Neg a \ is_Iiface a \ is_Oiface a)) (optimize_matches abstract_for_simple_firewall (upper_closure (iface_try_rewrite ipassmt rtblo (upper_closure (packet_assume_new rs)))))))" (*basically a copy&paste from transform_simple_fw_upper. but this one is way cleaner! refactor the other using this!*) theorem to_simple_firewall_without_interfaces: defines "newpkt p \ match_tcp_flags ipt_tcp_syn (p_tcp_flags p) \ p_tag_ctstate p = CT_New" assumes simplers: "simple_ruleset (rs:: 'i::len common_primitive rule list)" \ \well-formed ipassmt\ and wf_ipassmt1: "ipassmt_sanity_nowildcards (map_of ipassmt)" and wf_ipassmt2: "distinct (map fst ipassmt)" \ \There are no spoofed packets (probably by kernel's reverse path filter or our checker). This assumption implies that ipassmt lists ALL interfaces (!!).\ and nospoofing: "\(p::('i::len, 'a) tagged_packet_scheme). \ips. (map_of ipassmt) (Iface (p_iiface p)) = Some ips \ p_src p \ ipcidr_union_set (set ips)" \ \If a routing table was passed, the output interface for any packet we consider is decided based on it.\ and routing_decided: "\rtbl (p::('i,'a) tagged_packet_scheme). rtblo = Some rtbl \ output_iface (routing_table_semantics rtbl (p_dst p)) = p_oiface p" \ \A passed routing table is wellformed\ and correct_routing: "\rtbl. rtblo = Some rtbl \ correct_routing rtbl" \ \A passed routing table contains no interfaces with wildcard names\ and routing_no_wildcards: "\rtbl. rtblo = Some rtbl \ ipassmt_sanity_nowildcards (map_of (routing_ipassmt rtbl))" \ \the set of new packets, which are accepted is an overapproximations\ shows "{p::('i,'a) tagged_packet_scheme. (common_matcher, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p::('i,'a) tagged_packet_scheme. simple_fw (to_simple_firewall_without_interfaces ipassmt rtblo rs) p = Decision FinalAllow \ newpkt p}" and "\r \ set (to_simple_firewall_without_interfaces ipassmt rtblo rs). iiface (match_sel r) = ifaceAny \ oiface (match_sel r) = ifaceAny" proof - let ?rs1="packet_assume_new rs" let ?rs2="upper_closure ?rs1" let ?rs3="iface_try_rewrite ipassmt rtblo ?rs2" let ?rs4="upper_closure ?rs3" let ?rs5="optimize_matches abstract_for_simple_firewall ?rs4" let ?rs6="optimize_matches (abstract_primitive (\r. case r of Pos a \ is_Iiface a \ is_Oiface a | Neg a \ is_Iiface a \ is_Oiface a)) ?rs5" let ?rs7="upper_closure ?rs6" let ?\="(common_matcher, in_doubt_allow) :: ('i::len common_primitive, ('i, 'a) tagged_packet_scheme) match_tac" have "to_simple_firewall_without_interfaces ipassmt rtblo rs = to_simple_firewall ?rs7" by(simp add: to_simple_firewall_without_interfaces_def) from packet_assume_new_simple_ruleset[OF simplers] have s1: "simple_ruleset ?rs1" . from transform_upper_closure(2)[OF s1] have s2: "simple_ruleset ?rs2" . from iface_try_rewrite_simplers[OF s2] have s3: "simple_ruleset ?rs3" . from transform_upper_closure(2)[OF s3] have s4: "simple_ruleset ?rs4" . from optimize_matches_simple_ruleset[OF s4] have s5: "simple_ruleset ?rs5" . from optimize_matches_simple_ruleset[OF s5] have s6: "simple_ruleset ?rs6" . from transform_upper_closure(2)[OF s6] have s7: "simple_ruleset ?rs7" . from transform_upper_closure(3)[OF s1] have nnf2: "\r\set ?rs2. normalized_nnf_match (get_match r)" by simp from transform_upper_closure(3)[OF s3] have nnf4: "\r\set ?rs4. normalized_nnf_match (get_match r)" by simp have nnf5: "\r\set ?rs5. normalized_nnf_match (get_match r)" apply(intro optimize_matches_preserves) apply(simp add: abstract_for_simple_firewall_def) apply(rule abstract_primitive_preserves_normalized(5)) using nnf4 by(simp) have nnf6: "\r\set ?rs6. normalized_nnf_match (get_match r)" apply(intro optimize_matches_preserves) apply(rule abstract_primitive_preserves_normalized(5)) using nnf5 by(simp) from transform_upper_closure(3)[OF s6] have nnf7: "\r\set ?rs7. normalized_nnf_match (get_match r)" by simp (*subgoal @{term "check_simple_fw_preconditions ?rs7"}*) { fix m a assume r: "Rule m a \ set ?rs7" from s7 r have a: "(a = action.Accept \ a = action.Drop)" by(auto simp add: simple_ruleset_def) from abstract_for_simple_firewall_hasdisc have "\r\set ?rs5. \ has_disc is_CT_State (get_match r)" by(intro optimize_matches_preserves, blast) with abstract_primitive_preserves_nodisc[where disc'="is_CT_State"] have "\r\set ?rs6. \ has_disc is_CT_State (get_match r)" apply(intro optimize_matches_preserves) apply(simp) by blast with transform_upper_closure(4)[OF s6, where disc=is_CT_State] have "\r\set ?rs7. \ has_disc is_CT_State (get_match r)" by simp with r have no_CT: "\ has_disc is_CT_State m" by fastforce from abstract_for_simple_firewall_hasdisc have "\r\set ?rs5. \ has_disc is_L4_Flags (get_match r)" by(intro optimize_matches_preserves, blast) with abstract_primitive_preserves_nodisc[where disc'="is_L4_Flags"] have "\r\set ?rs6. \ has_disc is_L4_Flags (get_match r)" by(intro optimize_matches_preserves) auto with transform_upper_closure(4)[OF s6, where disc=is_L4_Flags] have "\r\set ?rs7. \ has_disc is_L4_Flags (get_match r)" by simp with r have no_L4_Flags: "\ has_disc is_L4_Flags m" by fastforce have "\r\set ?rs6. \ has_disc is_Iiface (get_match r)" by(intro optimize_matches_preserves abstract_primitive_nodisc) simp+ with transform_upper_closure(4)[OF s6, where disc=is_Iiface] have "\r\set ?rs7. \ has_disc is_Iiface (get_match r)" by simp with r have no_Iiface: "\ has_disc is_Iiface m" by fastforce have "\r\set ?rs6. \ has_disc is_Oiface (get_match r)" by(intro optimize_matches_preserves abstract_primitive_nodisc) simp+ with transform_upper_closure(4)[OF s6, where disc=is_Oiface] have "\r\set ?rs7. \ has_disc is_Oiface (get_match r)" by simp with r have no_Oiface: "\ has_disc is_Oiface m" by fastforce from no_Iiface no_Oiface have normalized_ifaces: "normalized_ifaces m" using has_disc_negated_disj_split has_disc_negated_has_disc normalized_ifaces_def by blast from transform_upper_closure(3)[OF s6] r have normalized: "normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m" by fastforce from transform_upper_closure(3)[OF s3, simplified] normalized_n_primitive_imp_not_disc_negated[OF wf_disc_sel_common_primitive(1)] normalized_n_primitive_imp_not_disc_negated[OF wf_disc_sel_common_primitive(2)] have "\r \ set ?rs4. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" apply(simp add: normalized_src_ports_def2 normalized_dst_ports_def2) by blast hence "\r \ set ?rs5. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" apply - apply(rule optimize_matches_preserves) apply(intro conjI) apply(intro abstract_for_simple_firewall_preserves_nodisc_negated, simp_all)+ by (simp add: abstract_for_simple_firewall_def abstract_primitive_preserves_nodisc) from this have no_ports_rs6: "\r \ set ?rs6. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" apply - apply(rule optimize_matches_preserves) apply(intro conjI) apply(intro abstract_primitive_preserves_nodisc_nedgated, simp_all)+ by (simp add: abstract_for_simple_firewall_def abstract_primitive_preserves_nodisc) from nnf4 abstract_for_simple_firewall_negated_ifaces_prots(2) have "\r\set ?rs5. \ has_disc_negated is_Prot False (get_match r)" by(intro optimize_matches_preserves) blast hence "\r\set ?rs6. \ has_disc_negated is_Prot False (get_match r)" by(intro optimize_matches_preserves abstract_primitive_preserves_nodisc_nedgated) simp+ with no_ports_rs6 have "\r\set ?rs7. \ has_disc_negated is_Prot False (get_match r)" by(intro transform_upper_closure(5)[OF s6]) simp+ with r have protocols: "normalized_protocols m" unfolding normalized_protocols_def by fastforce from no_CT no_L4_Flags normalized a normalized_ifaces protocols no_Iiface no_Oiface have "normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ normalized_ifaces m \ normalized_protocols m \ \ has_disc is_L4_Flags m \ \ has_disc is_CT_State m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m \ (a = action.Accept \ a = action.Drop)" and "\ has_disc is_Iiface m" and "\ has_disc is_Oiface m" apply - by(simp)+ (*fails due to is_MultiportPorts*) } hence simple_fw_preconditions: "check_simple_fw_preconditions ?rs7" and no_interfaces: "Rule m a \ set ?rs7 \ \ has_disc is_Iiface m \ \ has_disc is_Oiface m" for m a apply - subgoal unfolding check_simple_fw_preconditions_def by(clarify, rename_tac r, case_tac r, rename_tac m a, simp) by simp have "{p :: ('i,'a) tagged_packet_scheme. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p :: ('i,'a) tagged_packet_scheme. ?\,p\ \?rs1, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF s1]]) apply(subst approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]]) apply(rule Collect_cong) subgoal for p using packet_assume_new[where p = p] newpkt_def[where p = p] by auto done also have "{p. ?\,p\ \?rs1, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p. ?\,p\ \?rs2, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst transform_upper_closure(1)[OF s1]) by simp also have "\ = {p. ?\,p\ \?rs3, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(cases rtblo; simp; (subst iface_try_rewrite_rtbl[OF s2 nnf2] | subst iface_try_rewrite_no_rtbl[OF s2 nnf2])) using wf_ipassmt1 wf_ipassmt2 nospoofing wf_in_doubt_allow routing_no_wildcards correct_routing routing_decided by simp_all also have "\ = {p. ?\,p\ \?rs4, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst transform_upper_closure(1)[OF s3]) by simp finally have 1: "{p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p. ?\,p\ \?rs4, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" . from abstract_primitive_in_doubt_allow_generic(2)[OF primitive_matcher_generic_common_matcher nnf4 s4] have 2: "{p. ?\,p\ \?rs4, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \?rs5, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" by(auto simp add: abstract_for_simple_firewall_def) from abstract_primitive_in_doubt_allow_generic(2)[OF primitive_matcher_generic_common_matcher nnf5 s5] have 3: "{p. ?\,p\ \?rs5, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \?rs6, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" by(auto simp add: abstract_for_simple_firewall_def) have 4: "{p. ?\,p\ \?rs6, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} = {p. ?\,p\ \?rs7, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(subst transform_upper_closure(1)[OF s6]) by simp let ?fw="\rs p. approximating_bigstep_fun ?\ p rs Undecided" have approximating_rule: "\p. ?\,p\ \?rs7, Undecided\ \\<^sub>\ Decision FinalAllow \ ?fw ?rs7 p = Decision FinalAllow" using approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF s7]] by fast from 1 2 3 4 have "{p. ?\,p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. ?\,p\ \?rs7, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" by blast thus "{p. (common_matcher, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. simple_fw (to_simple_firewall_without_interfaces ipassmt rtblo rs) p = Decision FinalAllow \ newpkt p}" apply(safe) subgoal for p unfolding to_simple_firewall_without_interfaces_def using to_simple_firewall[OF simple_fw_preconditions, where p = p] approximating_rule[where p = p] by auto done (*the following proof to show that we don't have interfaces left is MADNESS*) have common_primitive_match_to_simple_match_nodisc: "Some sm = common_primitive_match_to_simple_match m' \ \ has_disc is_Iiface m' \ \ has_disc is_Oiface m' \ iiface sm = ifaceAny \ oiface sm = ifaceAny" if prems: "check_simple_fw_preconditions [Rule m' a']" for m' :: "'i common_primitive match_expr" and a' sm using prems proof(induction m' arbitrary: sm rule: common_primitive_match_to_simple_match.induct) case 18 thus ?case by(simp add: check_simple_fw_preconditions_def normalized_protocols_def) next case (13 m1 m2) thus ?case (*This is madness!!*) apply(simp add: check_simple_fw_preconditions_def) apply(case_tac "common_primitive_match_to_simple_match m1") apply(simp; fail) apply(case_tac "common_primitive_match_to_simple_match m2") apply(simp; fail) apply simp apply(rename_tac a aa) apply(case_tac a) apply(case_tac aa) apply(simp) apply(simp split: option.split_asm) using iface_conjunct_ifaceAny normalized_ifaces_def normalized_protocols_def by (metis has_disc_negated.simps(4) option.inject) qed(simp_all add: check_simple_fw_preconditions_def simple_match_any_def) have to_simple_firewall_no_ifaces: "(\m a. Rule m a \ set rs \ \ has_disc is_Iiface m \ \ has_disc is_Oiface m) \ \r\set (to_simple_firewall rs). iiface (match_sel r) = ifaceAny \ oiface (match_sel r) = ifaceAny" if pre1: "check_simple_fw_preconditions rs" for rs :: "'i common_primitive rule list" using pre1 apply(induction rs) apply(simp add: to_simple_firewall_simps; fail) apply simp apply(subgoal_tac "check_simple_fw_preconditions rs") prefer 2 subgoal by(simp add: check_simple_fw_preconditions_def) apply(rename_tac r rs, case_tac r) apply simp apply(simp add: to_simple_firewall_simps) apply(simp split: option.split) apply(intro conjI) apply blast apply(intro allI impI) apply(subgoal_tac "(\m\set (to_simple_firewall rs). iiface (match_sel m) = ifaceAny \ oiface (match_sel m) = ifaceAny)") prefer 2 subgoal by blast apply(simp) apply(rename_tac m' a' sm) apply(subgoal_tac " \ has_disc is_Iiface m' \ \ has_disc is_Oiface m'") prefer 2 subgoal by blast apply(subgoal_tac "check_simple_fw_preconditions [Rule m' a']") prefer 2 subgoal by(simp add: check_simple_fw_preconditions_def) apply(drule common_primitive_match_to_simple_match_nodisc) apply(simp_all) done from to_simple_firewall_no_ifaces[OF simple_fw_preconditions no_interfaces] show "\r \ set (to_simple_firewall_without_interfaces ipassmt rtblo rs). iiface (match_sel r) = ifaceAny \ oiface (match_sel r) = ifaceAny" unfolding to_simple_firewall_without_interfaces_def by(simp add: to_simple_firewall_def simple_fw_preconditions) qed end diff --git a/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy b/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy --- a/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy +++ b/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy @@ -1,1056 +1,1056 @@ theory LinuxRouter_OpenFlow_Translation imports IP_Addresses.CIDR_Split Automatic_Refinement.Misc (*TODO@Peter: rename and make available at better place :)*) Simple_Firewall.Generic_SimpleFw Semantics_OpenFlow OpenFlow_Matches OpenFlow_Action Routing.Linux_Router begin hide_const Misc.uncurry hide_fact Misc.uncurry_def (* For reference: iiface :: "iface" --"in-interface" oiface :: "iface" --"out-interface" src :: "(ipv4addr \ nat) " --"source IP address" dst :: "(ipv4addr \ nat) " --"destination" proto :: "protocol" sports :: "(16 word \ 16 word)" --"source-port first:last" dports :: "(16 word \ 16 word)" --"destination-port first:last" p_iiface :: string p_oiface :: string p_src :: ipv4addr p_dst :: ipv4addr p_proto :: primitive_protocol p_sport :: "16 word" p_dport :: "16 word" p_tcp_flags :: "tcp_flag set" p_payload :: string *) definition "route2match r = \iiface = ifaceAny, oiface = ifaceAny, src = (0,0), dst=(pfxm_prefix (routing_match r),pfxm_length (routing_match r)), - proto=ProtoAny, sports=(0,max_word), ports=(0,max_word)\" + proto=ProtoAny, sports=(0,- 1), ports=(0,- 1)\" definition toprefixmatch where "toprefixmatch m \ (let pm = PrefixMatch (fst m) (snd m) in if pm = PrefixMatch 0 0 then None else Some pm)" lemma prefix_match_semantics_simple_match: assumes some: "toprefixmatch m = Some pm" assumes vld: "valid_prefix pm" shows "prefix_match_semantics pm = simple_match_ip m" using some by(cases m) (clarsimp simp add: toprefixmatch_def ipset_from_cidr_def pfxm_mask_def fun_eq_iff prefix_match_semantics_ipset_from_netmask[OF vld] NOT_mask_shifted_lenword[symmetric] split: if_splits) definition simple_match_to_of_match_single :: "(32, 'a) simple_match_scheme \ char list option \ protocol \ (16 word \ 16 word) option \ (16 word \ 16 word) option \ of_match_field set" where "simple_match_to_of_match_single m iif prot sport dport \ uncurry L4Src ` option2set sport \ uncurry L4Dst ` option2set dport \ IPv4Proto ` (case prot of ProtoAny \ {} | Proto p \ {p}) \ \protocol is an 8 word option anyway...\ \ IngressPort ` option2set iif \ IPv4Src ` option2set (toprefixmatch (src m)) \ IPv4Dst ` option2set (toprefixmatch (dst m)) \ {EtherType 0x0800}" (* okay, we need to make sure that no packets are output on the interface they were input on. So for rules that don't have an input interface, we'd need to do a product over all interfaces, if we stay naive. The more smart way would be to insert a rule with the same match condition that additionally matches the input interface and drops. However, I'm afraid this is going to be very tricky to verify\ *) definition simple_match_to_of_match :: "32 simple_match \ string list \ of_match_field set list" where "simple_match_to_of_match m ifs \ (let - npm = (\p. fst p = 0 \ snd p = max_word); + npm = (\p. fst p = 0 \ snd p = - 1); sb = (\p. (if npm p then [None] else if fst p \ snd p then map (Some \ (\pfx. (pfxm_prefix pfx, NOT (pfxm_mask pfx)))) (wordinterval_CIDR_split_prefixmatch (WordInterval (fst p) (snd p))) else [])) in [simple_match_to_of_match_single m iif (proto m) sport dport. iif \ (if iiface m = ifaceAny then [None] else [Some i. i \ ifs, match_iface (iiface m) i]), sport \ sb (sports m), dport \ sb (dports m)] )" (* I wonder\ should I check whether list_all (match_iface (iiface m)) ifs instead of iiface m = ifaceAny? It would be pretty stupid if that wasn't the same, but you know\ *) lemma smtoms_eq_hlp: "simple_match_to_of_match_single r a b c d = simple_match_to_of_match_single r f g h i \ (a = f \ b = g \ c = h \ d = i)" (* In case this proof breaks: there are two alternate proofs in the repo. They are of similar quality, though. Good luck. *) proof(rule iffI,goal_cases) case 1 thus ?case proof(intro conjI) have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (IngressPort x)" by simp show "a = f" using 1 by(cases a; cases f) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Proto x\ \ P (IPv4Proto x)" by simp show "b = g" using 1 by(cases b; cases g) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (uncurry L4Src x)" by simp show "c = h" using 1 by(cases c; cases h) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (uncurry L4Dst x)" by simp show "d = i" using 1 by(cases d; cases i) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ qed qed simp lemma simple_match_to_of_match_generates_prereqs: "simple_match_valid m \ r \ set (simple_match_to_of_match m ifs) \ all_prerequisites r" unfolding simple_match_to_of_match_def Let_def proof(clarsimp, goal_cases) case (1 xiface xsrcp xdstp) note o = this show ?case unfolding simple_match_to_of_match_single_def all_prerequisites_def unfolding ball_Un proof((intro conjI; ((simp;fail)| - )), goal_cases) case 1 - have e: "(fst (sports m) = 0 \ snd (sports m) = max_word) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" + have e: "(fst (sports m) = 0 \ snd (sports m) = - 1) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" using o(1) unfolding simple_match_valid_alt Let_def by(clarsimp split: if_splits) show ?case using o(3) e by(elim disjE; simp add: option2set_def split: if_splits prod.splits uncurry_splits) next case 2 - have e: "(fst (dports m) = 0 \ snd (dports m) = max_word) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" + have e: "(fst (dports m) = 0 \ snd (dports m) = - 1) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" using o(1) unfolding simple_match_valid_alt Let_def by(clarsimp split: if_splits) show ?case using o(4) e by(elim disjE; simp add: option2set_def split: if_splits prod.splits uncurry_splits) qed qed lemma and_assoc: "a \ b \ c \ (a \ b) \ c" by simp lemmas custom_simpset = Let_def set_concat set_map map_map comp_def concat_map_maps set_maps UN_iff fun_app_def Set.image_iff abbreviation "simple_fw_prefix_to_wordinterval \ prefix_to_wordinterval \ uncurry PrefixMatch" lemma simple_match_port_alt: "simple_match_port m p \ p \ wordinterval_to_set (uncurry WordInterval m)" by(simp split: uncurry_splits) lemma simple_match_src_alt: "simple_match_valid r \ simple_match_ip (src r) p \ prefix_match_semantics (PrefixMatch (fst (src r)) (snd (src r))) p" by(cases "(src r)") (simp add: prefix_match_semantics_ipset_from_netmask2 prefix_to_wordset_ipset_from_cidr simple_match_valid_def valid_prefix_fw_def) lemma simple_match_dst_alt: "simple_match_valid r \ simple_match_ip (dst r) p \ prefix_match_semantics (PrefixMatch (fst (dst r)) (snd (dst r))) p" by(cases "(dst r)") (simp add: prefix_match_semantics_ipset_from_netmask2 prefix_to_wordset_ipset_from_cidr simple_match_valid_def valid_prefix_fw_def) lemma "x \ set (wordinterval_CIDR_split_prefixmatch w) \ valid_prefix x" using wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1] . lemma simple_match_to_of_matchI: assumes mv: "simple_match_valid r" assumes mm: "simple_matches r p" assumes ii: "p_iiface p \ set ifs" assumes ippkt: "p_l2type p = 0x800" shows eq: "\gr \ set (simple_match_to_of_match r ifs). OF_match_fields gr p = Some True" proof - - let ?npm = "\p. fst p = 0 \ snd p = max_word" + let ?npm = "\p. fst p = 0 \ snd p = - 1" let ?sb = "\p r. (if ?npm p then None else Some r)" obtain si where si: "case si of Some ssi \ p_sport p \ prefix_to_wordset ssi | None \ True" "case si of None \ True | Some ssi \ ssi \ set ( wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))" "si = None \ ?npm (sports r)" proof(cases "?npm (sports r)", goal_cases) case 1 (* True *) hence "(case None of None \ True | Some ssi \ p_sport p \ prefix_to_wordset ssi) \ (case None of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r))))" by simp with 1 show ?thesis by blast next case 2 (* False *) from mm have "p_sport p \ wordinterval_to_set (uncurry WordInterval (sports r))" by(simp only: simple_matches.simps simple_match_port_alt) then obtain ssi where ssi: "ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))" "p_sport p \ prefix_to_wordset ssi" using wordinterval_CIDR_split_existential by fast hence "(case Some ssi of None \ True | Some ssi \ p_sport p \ prefix_to_wordset ssi) \ (case Some ssi of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r))))" by simp with 2 show ?thesis by blast qed obtain di where di: "case di of Some ddi \ p_dport p \ prefix_to_wordset ddi | None \ True" "case di of None \ True | Some ddi \ ddi \ set ( wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))" "di = None \ ?npm (dports r)" proof(cases "?npm (dports r)", goal_cases) case 1 hence "(case None of None \ True | Some ssi \ p_dport p \ prefix_to_wordset ssi) \ (case None of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r))))" by simp with 1 show ?thesis by blast next case 2 from mm have "p_dport p \ wordinterval_to_set (uncurry WordInterval (dports r))" by(simp only: simple_matches.simps simple_match_port_alt) then obtain ddi where ddi: "ddi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))" "p_dport p \ prefix_to_wordset ddi" using wordinterval_CIDR_split_existential by fast hence "(case Some ddi of None \ True | Some ssi \ p_dport p \ prefix_to_wordset ssi) \ (case Some ddi of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r))))" by simp with 2 show ?thesis by blast qed show ?thesis proof let ?mf = "map_option (apsnd (wordNOT \ mask \ (-) 16) \ prefix_match_dtor)" let ?gr = "simple_match_to_of_match_single r (if iiface r = ifaceAny then None else Some (p_iiface p)) (if proto r = ProtoAny then ProtoAny else Proto (p_proto p)) (?mf si) (?mf di)" note mfu = simple_match_port.simps[of "fst (sports r)" "snd (sports r)", unfolded surjective_pairing[of "sports r",symmetric]] simple_match_port.simps[of "fst (dports r)" "snd (dports r)", unfolded surjective_pairing[of "dports r",symmetric]] note u = mm[unfolded simple_matches.simps mfu ord_class.atLeastAtMost_iff simple_packet_unext_def simple_packet.simps] note of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs] from u have ple: "fst (sports r) \ snd (sports r)" "fst (dports r) \ snd (dports r)" by force+ show eg: "?gr \ set (simple_match_to_of_match r ifs)" unfolding simple_match_to_of_match_def unfolding custom_simpset unfolding smtoms_eq_hlp proof(intro bexI, (intro conjI; ((rule refl)?)), goal_cases) case 2 thus ?case using ple(2) di apply(simp add: pfxm_mask_def prefix_match_dtor_def Set.image_iff split: option.splits prod.splits uncurry_splits) apply(erule bexI[rotated]) apply(simp split: prefix_match.splits) done next case 3 thus ?case using ple(1) si apply(simp add: pfxm_mask_def prefix_match_dtor_def Set.image_iff split: option.splits prod.splits uncurry_splits) apply(erule bexI[rotated]) apply(simp split: prefix_match.splits) done next case 4 thus ?case using u ii by(clarsimp simp: set_maps split: if_splits) next case 1 thus ?case using ii u by simp_all (metis match_proto.elims(2)) qed have dpm: "di = Some (PrefixMatch x1 x2) \ p_dport p && ~~ (mask (16 - x2)) = x1" for x1 x2 proof - have *: "di = Some (PrefixMatch x1 x2) \ prefix_match_semantics (the di) (p_dport p) \ p_dport p && ~~ (mask (16 - x2)) = x1" by(clarsimp simp: prefix_match_semantics_def pfxm_mask_def word_bw_comms;fail) have **: "pfx \ set (wordinterval_CIDR_split_prefixmatch ra) \ prefix_match_semantics pfx a = (a \ prefix_to_wordset pfx)" for pfx ra and a :: "16 word" by (fact prefix_match_semantics_wordset[OF wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1]]) have "\di = Some (PrefixMatch x1 x2); p_dport p \ prefix_to_wordset (PrefixMatch x1 x2); PrefixMatch x1 x2 \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))\ \ p_dport p && ~~ (mask (16 - x2)) = x1" using di(1,2) using * ** by auto thus "di = Some (PrefixMatch x1 x2) \ p_dport p && ~~ (mask (16 - x2)) = x1" using di(1,2) by auto qed have spm: "si = Some (PrefixMatch x1 x2) \ p_sport p && ~~ (mask (16 - x2)) = x1" for x1 x2 using si proof - have *: "si = Some (PrefixMatch x1 x2) \ prefix_match_semantics (the si) (p_sport p) \ p_sport p && ~~ (mask (16 - x2)) = x1" by(clarsimp simp: prefix_match_semantics_def pfxm_mask_def word_bw_comms;fail) have **: "pfx \ set (wordinterval_CIDR_split_prefixmatch ra) \ prefix_match_semantics pfx a = (a \ prefix_to_wordset pfx)" for pfx ra and a :: "16 word" by (fact prefix_match_semantics_wordset[OF wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1]]) have "\si = Some (PrefixMatch x1 x2); p_sport p \ prefix_to_wordset (PrefixMatch x1 x2); PrefixMatch x1 x2 \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))\ \ p_sport p && ~~ (mask (16 - x2)) = x1" using si(1,2) using * ** by auto thus "si = Some (PrefixMatch x1 x2) \ p_sport p && ~~ (mask (16 - x2)) = x1" using si(1,2) by auto qed show "OF_match_fields ?gr p = Some True" unfolding of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs[OF mv eg]] by(cases si; cases di) (simp_all add: simple_match_to_of_match_single_def OF_match_fields_unsafe_def spm option2set_def u ippkt prefix_match_dtor_def toprefixmatch_def dpm simple_match_dst_alt[OF mv, symmetric] simple_match_src_alt[OF mv, symmetric] split: prefix_match.splits) qed qed lemma prefix_match_00[simp,intro!]: "prefix_match_semantics (PrefixMatch 0 0) p" by (simp add: valid_prefix_def zero_prefix_match_all) lemma simple_match_to_of_matchD: assumes eg: "gr \ set (simple_match_to_of_match r ifs)" assumes mo: "OF_match_fields gr p = Some True" assumes me: "match_iface (oiface r) (p_oiface p)" assumes mv: "simple_match_valid r" shows "simple_matches r p" proof - from mv have validpfx: "valid_prefix (uncurry PrefixMatch (src r))" "valid_prefix (uncurry PrefixMatch (dst r))" "\pm. toprefixmatch (src r) = Some pm \ valid_prefix pm" "\pm. toprefixmatch (dst r) = Some pm \ valid_prefix pm" unfolding simple_match_valid_def valid_prefix_fw_def toprefixmatch_def by(simp_all split: uncurry_splits if_splits) from mo have mo: "OF_match_fields_unsafe gr p" unfolding of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs[OF mv eg]] by simp note this[unfolded OF_match_fields_unsafe_def] note eg[unfolded simple_match_to_of_match_def simple_match_to_of_match_single_def custom_simpset option2set_def] then guess x .. moreover from this(2) guess xa .. moreover from this(2) guess xb .. note xx = calculation(1,3) this { fix a b xc xa fix pp :: "16 word" have "\pp && ~~ (pfxm_mask xc) = pfxm_prefix xc\ \ prefix_match_semantics xc (pp)" for xc by(simp add: prefix_match_semantics_def word_bw_comms;fail) moreover have "pp \ wordinterval_to_set (WordInterval a b) \ a \ pp \ pp \ b" by simp moreover have "xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)) \ pp \ prefix_to_wordset xc \ pp \ wordinterval_to_set (WordInterval a b)" by(subst wordinterval_CIDR_split_prefixmatch) blast moreover have "\xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)); xa = Some (pfxm_prefix xc, ~~ (pfxm_mask xc)); prefix_match_semantics xc (pp)\ \ pp \ prefix_to_wordset xc" apply(subst(asm)(1) prefix_match_semantics_wordset) apply(erule wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1];fail) apply assumption done ultimately have "\xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)); xa = Some (pfxm_prefix xc, ~~ (pfxm_mask xc)); pp && ~~ (pfxm_mask xc) = pfxm_prefix xc\ \ a \ pp \ pp \ b" by metis } note l4port_logic = this show ?thesis unfolding simple_matches.simps proof(unfold and_assoc, (rule)+) show "match_iface (iiface r) (p_iiface p)" apply(cases "iiface r = ifaceAny") apply (simp add: match_ifaceAny) using xx(1) mo unfolding xx(4) OF_match_fields_unsafe_def apply(simp only: if_False set_maps UN_iff) apply(clarify) apply(rename_tac a; subgoal_tac "match_iface (iiface r) a") apply(clarsimp simp add: option2set_def;fail) apply(rule ccontr,simp;fail) done next show "match_iface (oiface r) (p_oiface p)" using me . next show "simple_match_ip (src r) (p_src p)" using mo unfolding xx(4) OF_match_fields_unsafe_def toprefixmatch_def by(clarsimp simp add: simple_packet_unext_def option2set_def validpfx simple_match_src_alt[OF mv] toprefixmatch_def split: if_splits) next show "simple_match_ip (dst r) (p_dst p)" using mo unfolding xx(4) OF_match_fields_unsafe_def toprefixmatch_def by(clarsimp simp add: simple_packet_unext_def option2set_def validpfx simple_match_dst_alt[OF mv] toprefixmatch_def split: if_splits) next show "match_proto (proto r) (p_proto p)" using mo unfolding xx(4) OF_match_fields_unsafe_def using xx(1) by(clarsimp simp add: singleton_iff simple_packet_unext_def option2set_def prefix_match_semantics_simple_match ball_Un split: if_splits protocol.splits) next show "simple_match_port (sports r) (p_sport p)" using mo xx(2) unfolding xx(4) OF_match_fields_unsafe_def by(cases "sports r") (clarsimp simp add: l4port_logic simple_packet_unext_def option2set_def prefix_match_semantics_simple_match split: if_splits) next show "simple_match_port (dports r) (p_dport p)" using mo xx(3) unfolding xx(4) OF_match_fields_unsafe_def by(cases "dports r") (clarsimp simp add: l4port_logic simple_packet_unext_def option2set_def prefix_match_semantics_simple_match split: if_splits) qed qed primrec annotate_rlen where "annotate_rlen [] = []" | "annotate_rlen (a#as) = (length as, a) # annotate_rlen as" lemma "annotate_rlen ''asdf'' = [(3, CHR ''a''), (2, CHR ''s''), (1, CHR ''d''), (0, CHR ''f'')]" by simp lemma fst_annotate_rlen_le: "(k, a) \ set (annotate_rlen l) \ k < length l" by(induction l arbitrary: k; simp; force) lemma distinct_fst_annotate_rlen: "distinct (map fst (annotate_rlen l))" using fst_annotate_rlen_le by(induction l) (simp, fastforce) lemma distinct_annotate_rlen: "distinct (annotate_rlen l)" using distinct_fst_annotate_rlen unfolding distinct_map by blast lemma in_annotate_rlen: "(a,x) \ set (annotate_rlen l) \ x \ set l" by(induction l) (simp_all, blast) lemma map_snd_annotate_rlen: "map snd (annotate_rlen l) = l" by(induction l) simp_all lemma "sorted_descending (map fst (annotate_rlen l))" by(induction l; clarsimp) (force dest: fst_annotate_rlen_le) lemma "annotate_rlen l = zip (rev [0.. *) primrec annotate_rlen_code where "annotate_rlen_code [] = (0,[])" | "annotate_rlen_code (a#as) = (case annotate_rlen_code as of (r,aas) \ (Suc r, (r, a) # aas))" lemma annotate_rlen_len: "fst (annotate_rlen_code r) = length r" by(induction r) (clarsimp split: prod.splits)+ lemma annotate_rlen_code[code]: "annotate_rlen s = snd (annotate_rlen_code s)" proof(induction s) case (Cons s ss) thus ?case using annotate_rlen_len[of ss] by(clarsimp split: prod.split) qed simp lemma suc2plus_inj_on: "inj_on (of_nat :: nat \ ('l :: len) word) {0..unat (max_word :: 'l word)}" proof(rule inj_onI) let ?mmw = "(max_word :: 'l word)" let ?mstp = "(of_nat :: nat \ 'l word)" fix x y :: nat assume "x \ {0..unat ?mmw}" "y \ {0..unat ?mmw}" hence se: "x \ unat ?mmw" "y \ unat ?mmw" by simp_all assume eq: "?mstp x = ?mstp y" note f = le_unat_uoi[OF se(1)] le_unat_uoi[OF se(2)] show "x = y" using eq le_unat_uoi se by metis qed lemma distinct_of_nat_list: (* TODO: Move to CaesarWordLemmaBucket *) "distinct l \ \e \ set l. e \ unat (max_word :: ('l::len) word) \ distinct (map (of_nat :: nat \ 'l word) l)" proof(induction l) let ?mmw = "(max_word :: 'l word)" let ?mstp = "(of_nat :: nat \ 'l word)" case (Cons a as) have "distinct as" "\e\set as. e \ unat ?mmw" using Cons.prems by simp_all note mIH = Cons.IH[OF this] moreover have "?mstp a \ ?mstp ` set as" proof have representable_set: "set as \ {0..unat ?mmw}" using \\e\set (a # as). e \ unat max_word\ by fastforce have a_reprbl: "a \ {0..unat ?mmw}" using \\e\set (a # as). e \ unat max_word\ by simp assume "?mstp a \ ?mstp ` set as" with inj_on_image_mem_iff[OF suc2plus_inj_on a_reprbl representable_set] have "a \ set as" by simp with \distinct (a # as)\ show False by simp qed ultimately show ?case by simp qed simp lemma annotate_first_le_hlp: "length l < unat (max_word :: ('l :: len) word) \ \e\set (map fst (annotate_rlen l)). e \ unat (max_word :: 'l word)" by(clarsimp) (meson fst_annotate_rlen_le less_trans nat_less_le) lemmas distinct_of_prio_hlp = distinct_of_nat_list[OF distinct_fst_annotate_rlen annotate_first_le_hlp] (* don't need these right now, but maybe later? *) lemma fst_annotate_rlen: "map fst (annotate_rlen l) = rev [0.. (of_nat :: nat \ ('l :: len) word)" assumes "length l \ unat (max_word :: 'l word)" shows "sorted_descending (map won (rev [0.. unat (max_word :: ('l :: len) word)" shows "sorted_descending (map fst (map (apfst (of_nat :: nat \ 'l word)) (annotate_rlen l)))" proof - let ?won = "(of_nat :: nat \ 'l word)" have "sorted_descending (map ?won (rev [0..l3 device to l2 forwarding\ definition "lr_of_tran_s3 ifs ard = ( [(p, b, case a of simple_action.Accept \ [Forward c] | simple_action.Drop \ []). (p,r,(c,a)) \ ard, b \ simple_match_to_of_match r ifs])" definition "oif_ne_iif_p1 ifs \ [(simple_match_any\oiface := Iface oif, iiface := Iface iif\, simple_action.Accept). oif \ ifs, iif \ ifs, oif \ iif]" definition "oif_ne_iif_p2 ifs = [(simple_match_any\oiface := Iface i, iiface := Iface i\, simple_action.Drop). i \ ifs]" definition "oif_ne_iif ifs = oif_ne_iif_p2 ifs @ oif_ne_iif_p1 ifs" (* order irrelephant *) (*value "oif_ne_iif [''a'', ''b'']"*) (* I first tried something like "oif_ne_iif ifs \ [(simple_match_any\oiface := Iface oi, iiface := Iface ii\, if oi = ii then simple_action.Drop else simple_action.Accept). oi \ ifs, ii \ ifs]", but making the statement I wanted with that was really tricky. Much easier to have the second element constant and do it separately. *) definition "lr_of_tran_s4 ard ifs \ generalized_fw_join ard (oif_ne_iif ifs)" definition "lr_of_tran_s1 rt = [(route2match r, output_iface (routing_action r)). r \ rt]" definition "lr_of_tran_fbs rt fw ifs \ let gfw = map simple_rule_dtor fw; \ \generalized simple fw, hopefully for FORWARD\ frt = lr_of_tran_s1 rt; \ \rt as fw\ prd = generalized_fw_join frt gfw in prd " definition "pack_OF_entries ifs ard \ (map (split3 OFEntry) (lr_of_tran_s3 ifs ard))" definition "no_oif_match \ list_all (\m. oiface (match_sel m) = ifaceAny)" definition "lr_of_tran rt fw ifs \ if \ (no_oif_match fw \ has_default_policy fw \ simple_fw_valid fw \ valid_prefixes rt \ has_default_route rt \ distinct ifs) then Inl ''Error in creating OpenFlow table: prerequisites not satisifed'' else ( let nrd = lr_of_tran_fbs rt fw ifs; ard = map (apfst of_nat) (annotate_rlen nrd) \ \give them a priority\ in - if length nrd < unat (max_word :: 16 word) + if length nrd < unat (- 1 :: 16 word) then Inr (pack_OF_entries ifs ard) else Inl ''Error in creating OpenFlow table: priority number space exhausted'') " definition "is_iface_name i \ i \ [] \ \Iface.iface_name_is_wildcard i" definition "is_iface_list ifs \ distinct ifs \ list_all is_iface_name ifs" lemma max_16_word_max[simp]: "(a :: 16 word) \ 0xffff" proof - have "0xFFFF = (- 1 :: 16 word)" by simp then show ?thesis by (simp only:) simp qed lemma replicate_FT_hlp: "x \ 16 \ y \ 16 \ replicate (16 - x) False @ replicate x True = replicate (16 - y) False @ replicate y True \ x = y" proof - let ?ns = "{0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15,16}" assume "x \ 16 \ y \ 16" hence "x \ ?ns" "y \ ?ns" by(simp; presburger)+ moreover assume "replicate (16 - x) False @ replicate x True = replicate (16 - y) False @ replicate y True" ultimately show "x = y" by simp (elim disjE; simp_all add: numeral_eq_Suc) (* that's only 289 subgoals after the elim *) qed lemma mask_inj_hlp1: "inj_on (mask :: nat \ 16 word) {0..16}" proof(intro inj_onI, goal_cases) case (1 x y) from 1(3) have oe: "of_bl (replicate (16 - x) False @ replicate x True) = (of_bl (replicate (16 - y) False @ replicate y True) :: 16 word)" unfolding mask_bl of_bl_rep_False . have "\z. z \ 16 \ length (replicate (16 - z) False @ replicate z True) = 16" by auto with 1(1,2) have ps: "replicate (16 - x) False @ replicate x True \ {bl. length bl = LENGTH(16)}" " replicate (16 - y) False @ replicate y True \ {bl. length bl = LENGTH(16)}" by simp_all from inj_onD[OF word_bl.Abs_inj_on, OF oe ps] show ?case using 1(1,2) by(fastforce intro: replicate_FT_hlp) qed lemma distinct_simple_match_to_of_match_portlist_hlp: fixes ps :: "(16 word \ 16 word)" shows "distinct ifs \ distinct (if fst ps = 0 \ snd ps = max_word then [None] else if fst ps \ snd ps then map (Some \ (\pfx. (pfxm_prefix pfx, ~~ (pfxm_mask pfx)))) (wordinterval_CIDR_split_prefixmatch (WordInterval (fst ps) (snd ps))) else [])" proof - assume di: "distinct ifs" { define wis where "wis = set (wordinterval_CIDR_split_prefixmatch (WordInterval (fst ps) (snd ps)))" fix x y :: "16 prefix_match" obtain xm xn ym yn where xyd[simp]: "x = PrefixMatch xm xn" "y = PrefixMatch ym yn" by(cases x; cases y) assume iw: "x \ wis" "y \ wis" and et: "(pfxm_prefix x, ~~ (pfxm_mask x)) = (pfxm_prefix y, ~~ (pfxm_mask y))" hence le16: "xn \ 16" "yn \ 16" unfolding wis_def using wordinterval_CIDR_split_prefixmatch_all_valid_Ball[unfolded Ball_def, THEN spec, THEN mp] by force+ with et have "16 - xn = 16 - yn" unfolding pfxm_mask_def by(auto intro: mask_inj_hlp1[THEN inj_onD]) hence "x = y" using et le16 using diff_diff_cancel by simp } note * = this show ?thesis apply(clarsimp simp add: smtoms_eq_hlp distinct_map wordinterval_CIDR_split_distinct) apply(subst comp_inj_on_iff[symmetric]; intro inj_onI) using * by simp_all qed lemma distinct_simple_match_to_of_match: "distinct ifs \ distinct (simple_match_to_of_match m ifs)" apply(unfold simple_match_to_of_match_def Let_def) apply(rule distinct_3lcomprI) subgoal by(induction ifs; clarsimp) subgoal by(fact distinct_simple_match_to_of_match_portlist_hlp) subgoal by(fact distinct_simple_match_to_of_match_portlist_hlp) subgoal by(simp_all add: smtoms_eq_hlp) done lemma inj_inj_on: "inj F \ inj_on F A" using subset_inj_on by auto (* TODO: include Word_Lib *) lemma no_overlaps_lroft_hlp2: "distinct (map fst amr) \ (\r. distinct (fm r)) \ distinct (concat (map (\(p, r, c, a). map (\b. (p, b, fs a c)) (fm r)) amr))" by(induction amr; force intro: injI inj_onI simp add: distinct_map split: prod.splits) lemma distinct_lroft_s3: "\distinct (map fst amr); distinct ifs\ \ distinct (lr_of_tran_s3 ifs amr)" unfolding lr_of_tran_s3_def by(erule no_overlaps_lroft_hlp2, simp add: distinct_simple_match_to_of_match) lemma no_overlaps_lroft_hlp3: "distinct (map fst amr) \ (aa, ab, ac) \ set (lr_of_tran_s3 ifs amr) \ (ba, bb, bc) \ set (lr_of_tran_s3 ifs amr) \ ac \ bc \ aa \ ba" apply(unfold lr_of_tran_s3_def) apply(clarsimp) apply(clarsimp split: simple_action.splits) apply(metis map_of_eq_Some_iff old.prod.inject option.inject) apply(metis map_of_eq_Some_iff old.prod.inject option.inject simple_action.distinct(2))+ done lemma no_overlaps_lroft_s3_hlp_hlp: (* I hlps *) "\distinct (map fst amr); OF_match_fields_unsafe ab p; ab \ ad \ ba \ bb; OF_match_fields_unsafe ad p; (ac, ab, ba) \ set (lr_of_tran_s3 ifs amr); (ac, ad, bb) \ set (lr_of_tran_s3 ifs amr)\ \ False" proof(elim disjE, goal_cases) case 1 have 4: "\distinct (map fst amr); (ac, ab, x1, x2) \ set amr; (ac, bb, x4, x5) \ set amr; ab \ bb\ \ False" for ab x1 x2 bb x4 x5 by (meson distinct_map_fstD old.prod.inject) have conjunctSomeProtoAnyD: "Some ProtoAny = simple_proto_conjunct a (Proto b) \ False" for a b using conjunctProtoD by force have 5: "\OF_match_fields_unsafe am p; OF_match_fields_unsafe bm p; am \ bm; am \ set (simple_match_to_of_match ab ifs); bm \ set (simple_match_to_of_match bb ifs); \ ab \ bb\ \ False" for ab bb am bm by(clarify | unfold simple_match_to_of_match_def smtoms_eq_hlp Let_def set_concat set_map de_Morgan_conj not_False_eq_True)+ (auto dest: conjunctSomeProtoAnyD cidrsplit_no_overlaps simp add: OF_match_fields_unsafe_def simple_match_to_of_match_single_def option2set_def comp_def split: if_splits cong: smtoms_eq_hlp) (*1min*) from 1 show ?case using 4 5 by(clarsimp simp add: lr_of_tran_s3_def) blast qed(metis no_overlaps_lroft_hlp3) lemma no_overlaps_lroft_s3_hlp: "distinct (map fst amr) \ distinct ifs \ no_overlaps OF_match_fields_unsafe (map (split3 OFEntry) (lr_of_tran_s3 ifs amr))" apply(rule no_overlapsI[rotated]) apply(subst distinct_map, rule conjI) subgoal by(erule (1) distinct_lroft_s3) subgoal apply(rule inj_inj_on) apply(rule injI) apply(rename_tac x y, case_tac x, case_tac y) apply(simp add: split3_def;fail) done subgoal apply(unfold check_no_overlap_def) apply(clarify) apply(unfold set_map) apply(clarify) apply(unfold split3_def prod.simps flow_entry_match.simps flow_entry_match.sel de_Morgan_conj) apply(clarsimp simp only:) apply(erule (1) no_overlaps_lroft_s3_hlp_hlp) apply simp apply assumption apply assumption apply simp done done lemma lr_of_tran_no_overlaps: assumes "distinct ifs" shows "Inr t = (lr_of_tran rt fw ifs) \ no_overlaps OF_match_fields_unsafe t" apply(unfold lr_of_tran_def Let_def pack_OF_entries_def) apply(simp split: if_splits) apply(thin_tac "t = _") apply(drule distinct_of_prio_hlp) apply(rule no_overlaps_lroft_s3_hlp[rotated]) subgoal by(simp add: assms) subgoal by(simp add: o_assoc) done lemma sorted_lr_of_tran_s3_hlp: "\x\set f. fst x \ a \ b \ set (lr_of_tran_s3 s f) \ fst b \ a" by(auto simp add: lr_of_tran_s3_def) lemma lr_of_tran_s3_Cons: "lr_of_tran_s3 ifs (a#ard) = ( [(p, b, case a of simple_action.Accept \ [Forward c] | simple_action.Drop \ []). (p,r,(c,a)) \ [a], b \ simple_match_to_of_match r ifs]) @ lr_of_tran_s3 ifs ard" by(clarsimp simp: lr_of_tran_s3_def) lemma sorted_lr_of_tran_s3: "sorted_descending (map fst f) \ sorted_descending (map fst (lr_of_tran_s3 s f))" apply(induction f) subgoal by(simp add: lr_of_tran_s3_def) apply(clarsimp simp: lr_of_tran_s3_Cons map_concat comp_def) apply(unfold sorted_descending_append) apply(simp add: sorted_descending_alt rev_map sorted_lr_of_tran_s3_hlp sorted_const) done lemma sorted_lr_of_tran_hlp: "(ofe_prio \ split3 OFEntry) = fst" by(simp add: fun_eq_iff comp_def split3_def) lemma lr_of_tran_sorted_descending: "Inr r = lr_of_tran rt fw ifs \ sorted_descending (map ofe_prio r)" apply(unfold lr_of_tran_def Let_def) apply(simp split: if_splits) apply(thin_tac "r = _") apply(unfold sorted_lr_of_tran_hlp pack_OF_entries_def split3_def[abs_def] fun_app_def map_map comp_def prod.case_distrib) apply(simp add: fst_def[symmetric]) apply(rule sorted_lr_of_tran_s3) apply(drule sorted_annotated[OF less_or_eq_imp_le, OF disjI1]) apply(simp add: o_assoc) done lemma lr_of_tran_s1_split: "lr_of_tran_s1 (a # rt) = (route2match a, output_iface (routing_action a)) # lr_of_tran_s1 rt" by(unfold lr_of_tran_s1_def list.map, rule) lemma route2match_correct: "valid_prefix (routing_match a) \ prefix_match_semantics (routing_match a) (p_dst p) \ simple_matches (route2match a) (p)" by(simp add: route2match_def simple_matches.simps match_ifaceAny match_iface_refl ipset_from_cidr_0 prefix_match_semantics_ipset_from_netmask2) lemma s1_correct: "valid_prefixes rt \ has_default_route (rt::('i::len) prefix_routing) \ \rm ra. generalized_sfw (lr_of_tran_s1 rt) p = Some (rm,ra) \ ra = output_iface (routing_table_semantics rt (p_dst p))" apply(induction rt) apply(simp;fail) apply(drule valid_prefixes_split) apply(clarsimp) apply(erule disjE) subgoal for a rt apply(case_tac a) apply(rename_tac routing_m metric routing_action) apply(case_tac routing_m) apply(simp add: valid_prefix_def pfxm_mask_def prefix_match_semantics_def generalized_sfw_def lr_of_tran_s1_def route2match_def simple_matches.simps match_ifaceAny match_iface_refl ipset_from_cidr_0 max_word_mask[where 'a = 'i, symmetric, simplified]) done subgoal apply(rule conjI) apply(simp add: generalized_sfw_def lr_of_tran_s1_def route2match_correct;fail) apply(simp add: route2match_def simple_matches.simps prefix_match_semantics_ipset_from_netmask2 lr_of_tran_s1_split generalized_sfw_simps) done done definition "to_OF_action a \ (case a of (p,d) \ (case d of simple_action.Accept \ [Forward p] | simple_action.Drop \ []))" definition "from_OF_action a = (case a of [] \ ('''',simple_action.Drop) | [Forward p] \ (p, simple_action.Accept))" lemma OF_match_linear_not_noD: "OF_match_linear \ oms p \ NoAction \ \ome. ome \ set oms \ \ (ofe_fields ome) p" apply(induction oms) apply(simp) apply(simp split: if_splits) apply blast+ done lemma s3_noaction_hlp: "\simple_match_valid ac; \simple_matches ac p; match_iface (oiface ac) (p_oiface p)\ \ OF_match_linear OF_match_fields_safe (map (\x. split3 OFEntry (x1, x, case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ [])) (simple_match_to_of_match ac ifs)) p = NoAction" apply(rule ccontr) apply(drule OF_match_linear_not_noD) apply(clarsimp) apply(rename_tac x) apply(subgoal_tac "all_prerequisites x") apply(drule simple_match_to_of_matchD) apply(simp add: split3_def) apply(subst(asm) of_match_fields_safe_eq2) apply(simp;fail)+ using simple_match_to_of_match_generates_prereqs by blast lemma aux: \v = Some x \ the v = x\ by simp lemma s3_correct: assumes vsfwm: "list_all simple_match_valid (map (fst \ snd) ard)" assumes ippkt: "p_l2type p = 0x800" assumes iiifs: "p_iiface p \ set ifs" assumes oiifs: "list_all (\m. oiface (fst (snd m)) = ifaceAny) ard" shows "OF_match_linear OF_match_fields_safe (pack_OF_entries ifs ard) p = Action ao \ (\r af. generalized_sfw (map snd ard) p = (Some (r,af)) \ (if snd af = simple_action.Drop then ao = [] else ao = [Forward (fst af)]))" unfolding pack_OF_entries_def lr_of_tran_s3_def fun_app_def using vsfwm oiifs apply(induction ard) subgoal by(simp add: generalized_sfw_simps) apply simp apply(clarsimp simp add: generalized_sfw_simps split: prod.splits) apply(intro conjI) (* make two subgoals from one *) subgoal for ard x1 ac ad ba apply(clarsimp simp add: OF_match_linear_append split: prod.splits) apply(drule simple_match_to_of_matchI[rotated]) apply(rule iiifs) apply(rule ippkt) apply blast apply(clarsimp simp add: comp_def) apply(drule OF_match_linear_match_allsameaction[where \=OF_match_fields_safe and pri = x1 and oms = "simple_match_to_of_match ac ifs" and act = "case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ []"]) apply(unfold OF_match_fields_safe_def comp_def) apply(erule aux) apply(clarsimp) apply(intro iffI) subgoal apply(rule exI[where x = ac]) apply(rule exI[where x = ad]) apply(rule exI[where x = ba]) apply(clarsimp simp: split3_def split: simple_action.splits flowtable_behavior.splits if_splits) done subgoal apply(clarsimp) apply(rename_tac b) apply(case_tac b) apply(simp_all) done done subgoal for ard x1 ac ad ba apply(simp add: OF_match_linear_append OF_match_fields_safe_def comp_def) apply(clarify) apply(subgoal_tac "OF_match_linear OF_match_fields_safe (map (\x. split3 OFEntry (x1, x, case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ [])) (simple_match_to_of_match ac ifs)) p = NoAction") apply(simp;fail) apply(erule (1) s3_noaction_hlp) apply(simp add: match_ifaceAny;fail) done done context notes valid_prefix_00[simp, intro!] begin lemma lr_of_tran_s1_valid: "valid_prefixes rt \ gsfw_valid (lr_of_tran_s1 rt)" unfolding lr_of_tran_s1_def route2match_def gsfw_valid_def list_all_iff apply(clarsimp simp: simple_match_valid_def valid_prefix_fw_def) apply(intro conjI) apply force apply(simp add: valid_prefixes_alt_def) done end lemma simple_match_valid_fbs_rlen: "\valid_prefixes rt; simple_fw_valid fw; (a, aa, ab, b) \ set (annotate_rlen (lr_of_tran_fbs rt fw ifs))\ \ simple_match_valid aa" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def] have "gsfw_valid (map simple_rule_dtor fw)" using gsfw_validI 1 by blast moreover have "gsfw_valid (lr_of_tran_s1 rt)" using 1 lr_of_tran_s1_valid by blast ultimately have "gsfw_valid (generalized_fw_join (lr_of_tran_s1 rt) (map simple_rule_dtor fw))" using gsfw_join_valid by blast moreover have "(aa, ab, b) \ set (lr_of_tran_fbs rt fw ifs)" using 1 using in_annotate_rlen by fast ultimately show ?thesis unfolding lr_of_tran_fbs_def Let_def gsfw_valid_def list_all_iff by fastforce qed lemma simple_match_valid_fbs: "\valid_prefixes rt; simple_fw_valid fw\ \ list_all simple_match_valid (map fst (lr_of_tran_fbs rt fw ifs))" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def] have "gsfw_valid (map simple_rule_dtor fw)" using gsfw_validI 1 by blast moreover have "gsfw_valid (lr_of_tran_s1 rt)" using 1 lr_of_tran_s1_valid by blast ultimately have "gsfw_valid (generalized_fw_join (lr_of_tran_s1 rt) (map simple_rule_dtor fw))" using gsfw_join_valid by blast thus ?thesis unfolding lr_of_tran_fbs_def Let_def gsfw_valid_def list_all_iff by fastforce qed lemma lr_of_tran_prereqs: "valid_prefixes rt \ simple_fw_valid fw \ lr_of_tran rt fw ifs = Inr oft \ list_all (all_prerequisites \ ofe_fields) oft" unfolding lr_of_tran_def pack_OF_entries_def lr_of_tran_s3_def Let_def apply(simp add: map_concat comp_def prod.case_distrib split3_def split: if_splits) apply(simp add: list_all_iff) apply(clarsimp) apply(drule simple_match_valid_fbs_rlen[rotated]) apply(simp add: list_all_iff;fail) apply(simp add: list_all_iff;fail) apply(rule simple_match_to_of_match_generates_prereqs; assumption) done (* TODO: move. where? *) lemma OF_unsafe_safe_match3_eq: " list_all (all_prerequisites \ ofe_fields) oft \ OF_priority_match OF_match_fields_unsafe oft = OF_priority_match OF_match_fields_safe oft" unfolding OF_priority_match_def[abs_def] proof(goal_cases) case 1 from 1 have "\packet. [f\oft . OF_match_fields_unsafe (ofe_fields f) packet] = [f\oft . OF_match_fields_safe (ofe_fields f) packet]" apply(clarsimp simp add: list_all_iff of_match_fields_safe_eq) using of_match_fields_safe_eq by(metis (mono_tags, lifting) filter_cong) thus ?case by metis qed lemma OF_unsafe_safe_match_linear_eq: " list_all (all_prerequisites \ ofe_fields) oft \ OF_match_linear OF_match_fields_unsafe oft = OF_match_linear OF_match_fields_safe oft" unfolding fun_eq_iff by(induction oft) (clarsimp simp add: list_all_iff of_match_fields_safe_eq)+ lemma simple_action_ne[simp]: "b \ simple_action.Accept \ b = simple_action.Drop" "b \ simple_action.Drop \ b = simple_action.Accept" using simple_action.exhaust by blast+ lemma map_snd_apfst: "map snd (map (apfst x) l) = map snd l" unfolding map_map comp_def snd_apfst .. lemma match_ifaceAny_eq: "oiface m = ifaceAny \ simple_matches m p = simple_matches m (p\p_oiface := any\)" by(cases m) (simp add: simple_matches.simps match_ifaceAny) lemma no_oif_matchD: "no_oif_match fw \ simple_fw fw p = simple_fw fw (p\p_oiface := any\)" by(induction fw) (auto simp add: no_oif_match_def simple_fw_alt dest: match_ifaceAny_eq) lemma lr_of_tran_fbs_acceptD: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" shows "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Accept) \ simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\)" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def, THEN generalized_fw_joinD] then guess r1 .. then guess r2 .. note r12 = this note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this from r12 rmra have oifra: "oif = ra" by simp from r12 have sfw: "simple_fw fw p = Decision FinalAllow" using simple_fw_iff_generalized_fw_accept by blast note ifupdateirrel = no_oif_matchD[OF s2, where any = " output_iface (routing_table_semantics rt (p_dst p))" and p = p, symmetric] show ?case unfolding simple_linux_router_nol12_def by(simp add: Let_def ifupdateirrel sfw oifra rmra split: Option.bind_splits option.splits) qed lemma lr_of_tran_fbs_acceptI: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" "has_default_policy fw" shows "simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ \r. generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Accept)" proof(goal_cases) from s2 have nud: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy state.distinct(1)) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1 have "simple_fw fw p = Decision FinalAllow" by(simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) then obtain r where r: "generalized_sfw (map simple_rule_dtor fw) p = Some (r, simple_action.Accept)" using simple_fw_iff_generalized_fw_accept by blast have oif_def: "oif = output_iface (routing_table_semantics rt (p_dst p))" using 1 by(cases p) (simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this show ?case unfolding lr_of_tran_fbs_def Let_def apply(rule exI) apply(rule generalized_fw_joinI) unfolding oif_def using rmra apply simp apply(rule r) done qed lemma lr_of_tran_fbs_dropD: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" shows "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop) \ simple_linux_router_nol12 rt fw p = None" proof(goal_cases) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1[unfolded lr_of_tran_fbs_def Let_def, THEN generalized_fw_joinD] obtain rr fr where "generalized_sfw (lr_of_tran_s1 rt) p = Some (rr, oif) \ generalized_sfw (map simple_rule_dtor fw) p = Some (fr, simple_action.Drop) \ Some r = simple_match_and rr fr" by presburger hence fd: "\u. simple_fw fw (p\p_oiface := u\) = Decision FinalDeny" unfolding ifupdateirrel using simple_fw_iff_generalized_fw_drop by blast show ?thesis by(clarsimp simp: simple_linux_router_nol12_def Let_def fd split: Option.bind_splits) qed lemma lr_of_tran_fbs_dropI: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" "has_default_policy fw" shows "simple_linux_router_nol12 rt fw p = None \ \r oif. generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" proof(goal_cases) from s2 have nud: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy state.distinct(1)) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1 have "simple_fw fw p = Decision FinalDeny" by(simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) then obtain r where r: "generalized_sfw (map simple_rule_dtor fw) p = Some (r, simple_action.Drop)" using simple_fw_iff_generalized_fw_drop by blast note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this show ?case unfolding lr_of_tran_fbs_def Let_def apply(rule exI) apply(rule exI[where x = ra]) apply(rule generalized_fw_joinI) using rmra apply simp apply(rule r) done qed lemma no_oif_match_fbs: "no_oif_match fw \ list_all (\m. oiface (fst (snd m)) = ifaceAny) (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))" proof(goal_cases) case 1 have c: "\mr ar mf af f a. \(mr, ar) \ set (lr_of_tran_s1 rt); (mf, af) \ simple_rule_dtor ` set fw; simple_match_and mr mf = Some a\ \ oiface a = ifaceAny" proof(goal_cases) case (1 mr ar mf af f a) have "oiface mr = ifaceAny" using 1(1) unfolding lr_of_tran_s1_def route2match_def by(clarsimp simp add: Set.image_iff) moreover have "oiface mf = ifaceAny" using 1(2) \no_oif_match fw\ unfolding no_oif_match_def simple_rule_dtor_def[abs_def] by(clarsimp simp: list_all_iff split: simple_rule.splits) fastforce ultimately show ?case using 1(3) by(cases a; cases mr; cases mf) (simp add: iface_conjunct_ifaceAny split: option.splits) qed have la: "list_all (\m. oiface (fst m) = ifaceAny) (lr_of_tran_fbs rt fw ifs)" unfolding lr_of_tran_fbs_def Let_def list_all_iff apply(clarify) apply(subst(asm) generalized_sfw_join_set) apply(clarsimp) using c by blast thus ?case proof(goal_cases) case 1 have *: "(\m. oiface (fst (snd m)) = ifaceAny) = (\m. oiface (fst m) = ifaceAny) \ snd" unfolding comp_def .. show ?case unfolding * list_all_map[symmetric] map_snd_apfst map_snd_annotate_rlen using la . qed qed lemma lr_of_tran_correct: fixes p :: "(32, 'a) simple_packet_ext_scheme" assumes nerr: "lr_of_tran rt fw ifs = Inr oft" and ippkt: "p_l2type p = 0x800" and ifvld: "p_iiface p \ set ifs" shows "OF_priority_match OF_match_fields_safe oft p = Action [Forward oif] \ simple_linux_router_nol12 rt fw p = (Some (p\p_oiface := oif\))" "OF_priority_match OF_match_fields_safe oft p = Action [] \ simple_linux_router_nol12 rt fw p = None" (* fun stuff: *) "OF_priority_match OF_match_fields_safe oft p \ NoAction" "OF_priority_match OF_match_fields_safe oft p \ Undefined" "OF_priority_match OF_match_fields_safe oft p = Action ls \ length ls \ 1" "\ls. length ls \ 1 \ OF_priority_match OF_match_fields_safe oft p = Action ls" proof - have s1: "valid_prefixes rt" "has_default_route rt" and s2: "has_default_policy fw" "simple_fw_valid fw" "no_oif_match fw" and difs: "distinct ifs" using nerr unfolding lr_of_tran_def by(simp_all split: if_splits) have "no_oif_match fw" using nerr unfolding lr_of_tran_def by(simp split: if_splits) note s2 = s2 this have unsafe_safe_eq: "OF_priority_match OF_match_fields_unsafe oft = OF_priority_match OF_match_fields_safe oft" "OF_match_linear OF_match_fields_unsafe oft = OF_match_linear OF_match_fields_safe oft" apply(subst OF_unsafe_safe_match3_eq; (rule lr_of_tran_prereqs s1 s2 nerr refl)+) apply(subst OF_unsafe_safe_match_linear_eq; (rule lr_of_tran_prereqs s1 s2 nerr refl)+) done have lin: "OF_priority_match OF_match_fields_safe oft = OF_match_linear OF_match_fields_safe oft" using OF_eq[OF lr_of_tran_no_overlaps lr_of_tran_sorted_descending, OF difs nerr[symmetric] nerr[symmetric]] unfolding fun_eq_iff unsafe_safe_eq by metis let ?ard = "map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs))" have oft_def: "oft = pack_OF_entries ifs ?ard" using nerr unfolding lr_of_tran_def Let_def by(simp split: if_splits) have vld: "list_all simple_match_valid (map (fst \ snd) ?ard)" unfolding fun_app_def map_map[symmetric] snd_apfst map_snd_apfst map_snd_annotate_rlen using simple_match_valid_fbs[OF s1(1) s2(2)] . have *: "list_all (\m. oiface (fst (snd m)) = ifaceAny) ?ard" using no_oif_match_fbs[OF s2(3)] . have not_undec: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy s2(1) state.simps(3)) have w1_1: "\oif. OF_match_linear OF_match_fields_safe oft p = Action [Forward oif] \ simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ oif = output_iface (routing_table_semantics rt (p_dst p))" proof(intro conjI, goal_cases) case (1 oif) note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD1, unfolded oft_def[symmetric], OF 1] hence "\r. generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, (oif, simple_action.Accept))" by(clarsimp split: if_splits) then obtain r where "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, (oif, simple_action.Accept))" unfolding map_map comp_def snd_apfst map_snd_annotate_rlen by blast thus ?case using lr_of_tran_fbs_acceptD[OF s1 s2(3)] by metis thus "oif = output_iface (routing_table_semantics rt (p_dst p))" by(cases p) (clarsimp simp: simple_linux_router_nol12_def Let_def not_undec split: Option.bind_splits state.splits final_decision.splits) qed have w1_2: "\oif. simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ OF_match_linear OF_match_fields_safe oft p = Action [Forward oif]" proof(goal_cases) case (1 oif) note lr_of_tran_fbs_acceptI[OF s1 s2(3) s2(1) this, of ifs] then guess r .. note r = this hence "generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, (oif, simple_action.Accept))" unfolding map_snd_apfst map_snd_annotate_rlen . moreover note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD2, unfolded oft_def[symmetric], of "[Forward oif]"] ultimately show ?case by simp qed show w1: "\oif. (OF_priority_match OF_match_fields_safe oft p = Action [Forward oif]) = (simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\))" unfolding lin using w1_1 w1_2 by blast show w2: "(OF_priority_match OF_match_fields_safe oft p = Action []) = (simple_linux_router_nol12 rt fw p = None)" unfolding lin proof(rule iffI, goal_cases) case 1 note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD1, unfolded oft_def[symmetric], OF 1] then obtain r oif where roif: "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" unfolding map_snd_apfst map_snd_annotate_rlen by(clarsimp split: if_splits) note lr_of_tran_fbs_dropD[OF s1 s2(3) this] thus ?case . next case 2 note lr_of_tran_fbs_dropI[OF s1 s2(3) s2(1) this, of ifs] then obtain r oif where "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" by blast hence "generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, oif, simple_action.Drop)" unfolding map_snd_apfst map_snd_annotate_rlen . moreover note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD2, unfolded oft_def[symmetric], of "[]"] ultimately show ?case by force qed have lr_determ: "\a. simple_linux_router_nol12 rt fw p = Some a \ a = p\p_oiface := output_iface (routing_table_semantics rt (p_dst p))\" by(clarsimp simp: simple_linux_router_nol12_def Let_def not_undec split: Option.bind_splits state.splits final_decision.splits) show notno: "OF_priority_match OF_match_fields_safe oft p \ NoAction" apply(cases "simple_linux_router_nol12 rt fw p") using w2 apply(simp) using w1[of "output_iface (routing_table_semantics rt (p_dst p))"] apply(simp) apply(drule lr_determ) apply(simp) done show notub: "OF_priority_match OF_match_fields_safe oft p \ Undefined" unfolding lin using OF_match_linear_ne_Undefined . show notmult: "\ls. OF_priority_match OF_match_fields_safe oft p = Action ls \ length ls \ 1" apply(cases "simple_linux_router_nol12 rt fw p") using w2 apply(simp) using w1[of "output_iface (routing_table_semantics rt (p_dst p))"] apply(simp) apply(drule lr_determ) apply(clarsimp) done show "\ls. length ls \ 1 \ OF_priority_match OF_match_fields_safe oft p = Action ls" apply(cases "OF_priority_match OF_match_fields_safe oft p") using notmult apply blast using notno apply blast using notub apply blast done qed end diff --git a/thys/LOFT/OpenFlow_Documentation.thy b/thys/LOFT/OpenFlow_Documentation.thy --- a/thys/LOFT/OpenFlow_Documentation.thy +++ b/thys/LOFT/OpenFlow_Documentation.thy @@ -1,533 +1,533 @@ text_raw\ \twocolumn \columnsep 2pc % Space between columns \textwidth 42pc % Width of text line. \part{Documentation} \label{part2} \ section\Configuration Translation\ text_raw\\label{sec:conv}\ text\ All the results we present in this section are formalized and verified in Isabelle/HOL~\cite{nipkow2002isabelle}. This means that their formal correctness can be trusted a level close to absolute certainty. The definitions and lemmas stated here are merely a repetition of lemmas stated in other theory files. This means that they have been directly set to this document from Isabelle and no typos or hidden assumptions are possible. Additionally, it allows us to omit various helper lemmas that do not help the understanding. However, it causes some notation inaccuracy, as type and function definitions are stated as lemmas or schematic goals. \ theory OpenFlow_Documentation (*<*) imports LinuxRouter_OpenFlow_Translation Featherweight_OpenFlow_Comparison "HOL-Library.LaTeXsugar" begin (*>*) subsection\Linux Firewall Model\ text_raw\\label{sec:lfw}\ text\We want to write a program that translates the configuration of a linux firewall to that of an OpenFlow switch. We furthermore want to verify that translation. For this purpose, we need a clear definition of the behavior of the two device types -- we need their models and semantics. In case of a linux firewall, this is problematic because a linux firewall is a highly complex device that is ultimately capable of general purpose computation. Creating a comprehensive semantics that encompasses all possible configuration types of a linux firewall is thus highly non-trivial and not useful for the purpose of analysis. We decided to approach the problem from the other side: we created a model that includes only the most basic features. (This implies neglecting IPv6.) Fortunately, many of the highly complex features are rarely essential and even our basic model is still of some use. \ text\We first divided the firewall into subsystems. Given a routing table @{term rt}, the firewall rules @{term fw}, the routing decision for a packet @{term p} can be obtained by @{term "routing_table_semantics rt (p_dst p)"}, the firewall decision by @{term "simple_fw fw p"}. We draft the first description of our linux router model: \begin{enumerate} \item The destination MAC address of an arriving packet is checked: Does it match the MAC address of the ingress port? If it does, we continue, otherwise, the packet is discarded. \item The routing decision @{term "rd \ routing_table_semantics rt p"} is obtained. \item The packet's output interface is updated based on @{term rd}\footnote{Note that we assume a packet model with input and output interfaces. The origin of this is explained in Section~\ref{sec:lfwfw}}. \item The firewall is queried for a decision: @{term "simple_fw fw p"}. If the decision is to @{const[names_short] simple_action.Drop}, the packet is discarded. \item The next hop is computed: If @{term rd} provides a next hop, that is used. Otherwise, the destination address of the packet is used. \item The MAC address of the next hop is looked up; the packet is updated with it and sent. \end{enumerate} We decided that this description is best formalized as an abortable program in the option monad:\ lemma "simple_linux_router rt fw mlf ifl p \ do { _ \ iface_packet_check ifl p; let rd \ \(routing decision)\ = routing_table_semantics rt (p_dst p); let p = p\p_oiface := output_iface rd\; let fd \ \(firewall decision)\ = simple_fw fw p; _ \ (case fd of Decision FinalAllow \ Some () | Decision FinalDeny \ None); let nh = (case next_hop rd of None \ p_dst p | Some a \ a); ma \ mlf nh; Some (p\p_l2dst := ma\) }" unfolding fromMaybe_def[symmetric] by(fact simple_linux_router_def) text\where @{term "mlf :: ipv4addr \ 48 word"} is a function that looks up the MAC address for an IP address.\ text\There are already a few important aspects that have not been modelled, but they are not core essential for the functionality of a firewall. Namely, there is no local traffic from/to the firewall. This is problematic since this model can not generate ARP replies --- thus, an equivalent OpenFlow device will not do so, either. Furthermore, this model is problematic because it requires access to a function that looks up a MAC address, something that may not be known at the time of time running a translation to an OpenFlow configuration. \ text\It is possible to circumvent these problems by inserting static ARP table entries in the directly connected devices and looking up their MAC addresses \emph{a priori}. A test-wise implementation of the translation based on this model showed acceptable results. However, we deemed the \emph{a priori} lookup of the MAC addresses to be rather inelegant and built a second model.\ definition "simple_linux_router_altered rt fw ifl p \ do { let rd = routing_table_semantics rt (p_dst p); let p = p\p_oiface := output_iface rd\; _ \ if p_oiface p = p_iiface p then None else Some (); let fd = simple_fw fw p; _ \ (case fd of Decision FinalAllow \ Some () | Decision FinalDeny \ None); Some p }" (* TODO: Would a router actually forward a packet on the same interface? *) text\In this model, all access to the MAC layer has been eliminated. This is done by the approximation that the firewall will be asked to route a packet (i.e. be addressed on the MAC layer) iff the destination IP address of the packet causes it to be routed out on a different interface. Because this model does not insert destination MAC addresses, the destination MAC address has to be already correct when the packet is sent. This can only be achieved by changing the subnet of all connected device, moving them into one common subnet\footnote{There are cases where this is not possible --- A limitation of our system.}. \ text\ While a test-wise implementation based on this model also showed acceptable results, the model is still problematic. The check @{term "p_oiface p = p_iiface p"} and the firewall require access to the output interface. The details of why this cannot be provided are be elaborated in Section~\ref{sec:convi}. The intuitive explanation is that an OpenFlow match can not have a field for the output interface. We thus simplified the model even further: \ lemma "simple_linux_router_nol12 rt fw p \ do { let rd = routing_table_semantics rt (p_dst p); let p = p\p_oiface := output_iface rd\; let fd = simple_fw fw p; _ \ (case fd of Decision FinalAllow \ Some () | Decision FinalDeny \ None); Some p }" by(fact simple_linux_router_nol12_def) text\We continue with this definition as a basis for our translation. Even this strongly altered version and the original linux firewall still behave the same in a substantial amount of cases:\ theorem "\iface_packet_check ifl pii \ None; mlf (case next_hop (routing_table_semantics rt (p_dst pii)) of None \ p_dst pii | Some a \ a) \ None\ \ \x. map_option (\p. p\p_l2dst := x\) (simple_linux_router_nol12 rt fw pii) = simple_linux_router rt fw mlf ifl pii" by(fact rtr_nomac_eq[unfolded fromMaybe_def]) text\The conditions are to be read as ``The check whether a received packet has the correct destination MAC never returns @{const False}'' and ``The next hop MAC address for all packets can be looked up''. Obviously, these conditions do not hold for all packets. We will show an example where this makes a difference in Section~\ref{sec:mnex}.\ subsubsection\Routing Table\ text_raw\\label{sec:lfwr}\ text\ The routing system in linux features multiple tables and a system that can use the iptables firewall and an additional match language to select a routing table. Based on our directive, we only focused on the single most used \texttt{main} routing table.\ text\ We define a routing table entry to be a record (named tuple) of a prefix match, a metric and the routing action, which in turn is a record of an output interface and an optional next-hop address.\ schematic_goal "(?rtbl_entry :: ('a::len) routing_rule) = \ routing_match = PrefixMatch pfx len, metric = met, routing_action = \ output_iface = oif_string, next_hop = (h :: 'a word option) \ \" .. text\A routing table is then a list of these entries:\ lemma "(rtbl :: ('a :: len) prefix_routing) = (rtbl :: 'a routing_rule list)" by rule text\Not all members of the type @{type prefix_routing} are sane routing tables. There are three different validity criteria that we require so that our definitions are adequate. \begin{itemize} \item The prefixes have to be 0 in bits exceeding their length. \item There has to be a default rule, i.e. one with prefix length 0. With the condition above, that implies that all its prefix bits are zero and it thus matches any address. \item The entries have to be sorted by prefix length and metric. \end{itemize} The first two are set into code in the following way: \ lemma "valid_prefix (PrefixMatch pfx len) \ pfx && (2 ^ (32 - len) - 1) = (0 :: 32 word)" by (simp add: valid_prefix_def pfxm_mask_def mask_eq_decr_exp and.commute) lemma "has_default_route rt \ (\r \ set rt. pfxm_length (routing_match r) = 0)" by(fact has_default_route_alt) text\The third is not needed in any of the further proofs, so we omit it.\ text\The semantics of a routing table is to simply traverse the list until a matching entry is found.\ schematic_goal "routing_table_semantics (rt_entry # rt) dst_addr = (if prefix_match_semantics (routing_match rt_entry) dst_addr then routing_action rt_entry else routing_table_semantics rt dst_addr)" by(fact routing_table_semantics.simps) text\If no matching entry is found, the behavior is undefined.\ subsubsection\iptables Firewall\ text_raw\\label{sec:lfwfw}\ text\The firewall subsystem in a linux router is not any less complex than any of the of the other systems. Fortunately, this complexity has been dealt with in~\cite{diekmann2016verified,Iptables_Semantics-AFP} already and we can directly use the result.\ text\In short, one of the results is that a complex \emph{iptables} configuration can be simplified to be represented by a single list of matches that only support the following match conditions: \begin{itemize} \item (String) prefix matches on the input and output interfaces. \item A @{type prefix_match} on the source and destination IP address. \item An exact match on the layer 4 protocol. \item Interval matches on the source or destination port, e.g. @{term "p\<^sub>d \ {(1::16 word)..1023}"} \end{itemize} The model/type of the packet is adjusted to fit that: it is a record of the fields matched on. This also means that input and output interface are coded to the packet. Given that this information is usually stored alongside the packet content, this can be deemed a reasonable model. In case the output interface is not needed (e.g., when evaluating an OpenFlow table), it can simply be left blank. Obviously, a simplification into the above match type cannot always produce an equivalent firewall, and the set of accepted packets has to be over- or underapproximated. The reader interested in the details of this is strongly referred to~\cite{diekmann2016verified}; we are simply going to continue with the result: @{const simple_fw}. \ text\One property of the simplification is worth noting here: The simplified firewall does not know state and the simplification approximates stateful matches by stateless ones. Thus, the overapproximation of a stateful firewall ruleset that begins with accepting packets of established connections usually begins with a rule that accepts all packets. Dealing with this by writing a meaningful simplification of stateful firewalls is future work. \ subsection\OpenFlow Switch Model\ text\In this section, we present our model of an OpenFlow switch. The requirements for this model are derived from the fact that it models devices that are the target of a configuration translation. This has two implications: \begin{itemize} \item All configurations that are representable in our model should produce the correct behavior wrt. their semantics. The problem is that correct here means that the behavior is the same that any real device would produce. Since we cannot possibly account for all device types, we instead focus on those that conform to the OpenFlow specifications. To account for the multiple different versions of the specification (e.g.~\cite{specification10,specification15}), we tried making our model a subset of both the oldest stable version 1.0~\cite{specification10} and the newest available specification version 1.5.1~\cite{specification15}. \item Conversely, our model does not need to represent all possible behavior of an OpenFlow switch, just the behavior that can be invoked by the result of our translation. This is especially useful regarding for controller interaction, but also for MPLS or VLANs, which we did not model in Section \ref{sec:lfw}. \end{itemize}\ text\More concretely, we set the following rough outline for our model. \begin{itemize} \item A switch consists of a single flow table. \item A flow table entry consists of a priority, a match condition and an action list. \item The only possible action (we require) is to forward the packet on a port. \item We do not model controller interaction. \end{itemize} Additionally, we decided that we wanted to be able to ensure the validity of the flow table in all qualities, i.e. we want to model the conditions `no overlapping flow entries appear', `all match conditions have their necessary preconditions'. The details of this are explained in the following sections. \ subsubsection\Matching Flow Table entries\ text_raw\\label{sec:of_match}\ text\Table 3 of Section 3.1 of \cite{specification10} gives a list of required packet fields that can be used to match packets. This directly translates into the type for a match expression on a single field:\ schematic_goal "(field_match :: of_match_field) \ { IngressPort (?s::string), EtherSrc (?as::48 word), EtherDst (?ad::48 word), EtherType (?t::16 word), VlanId (?i::16 word), VlanPriority (?p::16 word), IPv4Src (?pms::32 prefix_match), IPv4Dst (?pmd::32 prefix_match), IPv4Proto (?ipp :: 8 word), L4Src (?ps :: 16 word) (?ms :: 16 word), L4Dst (?pd :: 16 word) (?md :: 16 word) }" by(fact of_match_field_typeset) text\ Two things are worth additional mention: L3 and L4 ``addressess''. The @{term IPv4Src} and @{term IPv4Dst} matches are specified as ``can be subnet masked'' in~\cite{specification10}, whereras~\cite{specification15} states clearly that arbitrary bitmasks can be used. We took the conservative approach here. Our alteration of @{term L4Src} and @{term L4Dst} is more grave. While~\cite{specification10} does not state anything about layer 4 ports and masks, \cite{specification15} specifically forbids using masks on them. Nevertheless, OpenVSwitch \cite{openvswitch} and some other implementations support them. We will explain in detail why we must include bitmasks on layer 4 ports to obtain a meaningful translation in Section~\ref{sec:convi}.\ text\One @{type of_match_field} is not enough to classify a packet. To match packets, we thus use entire sets of match fields. As Guha \emph{et al.}~\cite{guha2013machine} noted\footnote{See also: \cite[\
2.3]{michaelis2016middlebox}}, executing a set of given @{type of_match_field}s on a packet requires careful consideration. For example, it is not meaningful to use @{term IPv4Dst} if the given packet is not actually an IP packet, i.e. @{term IPv4Dst} has the prerequisite of @{term "EtherType 0x0800"} being among the match fields. Guha \emph{et al.} decided to use the fact that the preconditions can be arranged on a directed acyclic graph (or rather: an acyclic forest). They evaluated match conditions in a manner following that graph: first, all field matches without preconditions are evaluated. Upon evaluating a field match (e.g., @{term "EtherType 0x0800"}), the matches that had their precondition fulfilled by it (e.g., @{term IPv4Src} and @{term IPv4Src} in this example) are evalutated. This mirrors the faulty behavior of some implementations (see \cite{guha2013machine}). Adopting that behavior into our model would mean that any packet matches against the field match set @{term "{IPv4Dst (PrefixMatch 134744072 32)}"} instead of just those destined for 8.8.8.8 or causing an error. We found this to be unsatisfactory.\ text\To solve this problem, we made three definitions. The first, @{term match_no_prereq} matches an @{type of_match_field} against a packet without considering prerequisites. The second, @{term prerequisites}, checks for a given @{type of_match_field} whether its prerequisites are in a set of given match fields. Especially: \ lemma "prerequisites (VlanPriority pri) m = (\id. let v = VlanId id in v \ m \ prerequisites v m)" "prerequisites (IPv4Proto pr) m = (let v = EtherType 0x0800 in v \ m \ prerequisites v m)" "prerequisites (IPv4Src a) m = (let v = EtherType 0x0800 in v \ m \ prerequisites v m)" "prerequisites (IPv4Dst a) m = (let v = EtherType 0x0800 in v \ m \ prerequisites v m)" "prerequisites (L4Src p msk) m = (\proto \ {TCP,UDP,L4_Protocol.SCTP}. let v = IPv4Proto proto in v \ m \ prerequisites v m)" "prerequisites (L4Dst p msk) m = prerequisites (L4Src undefined undefined) m" by(fact prerequisites.simps)+ text\Then, to actually match a set of @{type of_match_field} against a packet, we use the option type:\ lemma "OF_match_fields m p = (if \f \ m. \prerequisites f m then None else if \f \ m. match_no_prereq f p then Some True else Some False)" by(fact OF_match_fields_alt) subsubsection\Evaluating a Flow Table\ text\In the previous section, we explained how we match the set of match fields belonging to a single flow entry against a packet. This section explains how the correct flow entry from a table can be selected. To prevent to much entanglement with the previous section, we assume an arbitrary match function @{term "\ :: 'match_field set \ 'packet \ bool"}. This function @{term "\"} takes the match condition @{term m} from a flow entry @{term "OFEntry (priority::16 word) (m::'match_field set) action"} and decides whether a packet matches those.\ text\The flow table is simply a list of flow table entries @{type flow_entry_match}. Deciding the right flow entry to use for a given packet is explained in the OpenFlow specification \cite{specification10}, Section 3.4: \begin{quote} Packets are matched against flow entries based on prioritization. An entry that specifies an exact match (i.e., has no wildcards) is always the highest priority\footnote{This behavior has been deprecated.}. All wildcard entries have a priority associated with them. Higher priority entries must match before lower priority ones. If multiple entries have the same priority, the switch is free to choose any ordering. \end{quote} We use the term ``overlapping'' for the flow entries that can cause a packet to match multiple flow entries with the same priority. Guha \emph{et al.}~\cite{guha2013machine} have dealt with overlapping. However, the semantics for a flow table they presented \cite[Figure 5]{guha2013machine} is slightly different from what they actually used in their theory files. We have tried to reproduce the original inductive definition (while keeping our abstraction @{term \}), in Isabelle/HOL\footnote{The original is written in Coq~\cite{barras1997coq} and we can not use it directly.}:\ lemma "\ (ofe_fields fe) p = True \ \fe' \ set (ft1 @ ft2). ofe_prio fe' > ofe_prio fe \ \ (ofe_fields fe') p = False \ guha_table_semantics \ (ft1 @ fe # ft2) p (Some (ofe_action fe))" "\fe \ set ft. \ (ofe_fields fe) p = False \ guha_table_semantics \ ft p None" by(fact guha_matched guha_unmatched)+ text\Guha \emph{et al.} have deliberately made their semantics non-deterministic, to match the fact that the switch ``may choose any ordering''. This can lead to undesired results:\ lemma "CARD('action) \ 2 \ \ff. \ ff p \ \ft (a1 :: 'action) (a2 :: 'action). a1 \ a2 \ guha_table_semantics \ ft p (Some a1) \ guha_table_semantics \ ft p (Some a2)" by(fact guha_table_semantics_ex2res) text\This means that, given at least two distinct actions exist and our matcher @{term \} is not false for all possible match conditions, we can say that a flow table and two actions exist such that both actions are executed. This can be misleading, as the switch might choose an ordering on some flow table and never execute some of the (overlapped) actions.\ text\Instead, we decided to follow Section 5.3 of the specification \cite{specification15}, which states: \begin{quote} If there are multiple matching flow entries, the selected flow entry is explicitly undefined. \end{quote} This still leaves some room for interpretation, but it clearly states that overlapping flow entries are undefined behavior, and undefined behavior should not be invoked. Thus, we came up with a semantics that clearly indicates when undefined behavior has been invoked:\ lemma "OF_priority_match \ flow_entries packet = ( let m = filter (\f. \ (ofe_fields f) packet) flow_entries; m' = filter (\f. \fo \ set m. ofe_prio fo \ ofe_prio f) m in case m' of [] \ NoAction | [s] \ Action (ofe_action s) | _ \ Undefined)" unfolding OF_priority_match_def .. text\The definition works the following way\footnote{Note that the order of the flow table entries is irrelevant. We could have made this definition on sets but chose not to for consistency.}: \begin{enumerate} \item The flow table is filtered for those entries that match, the result is called $m$. \item $m$ is filtered again, leaving only those entries for which no entries with lower priority could be found, i.e. the matching flow table entries with minimal priority. The result is called $m'$. \item A case distinction on $m'$ is made. If only one matching entry was found, its action is returned for execution. If $m$ is empty, the flow table semantics returns @{term NoAction} to indicate that the flow table does not decide an action for the packet. If, not zero or one entry is found, but more, the special value @{term Undefined} for indicating undefined behavior is returned. \end{enumerate} The use of @{term Undefined} immediately raises the question in which condition it cannot occur. We give the following definition:\ lemma "check_no_overlap \ ft = (\a \ set ft. \b \ set ft. (a \ b \ ofe_prio a = ofe_prio b) \ \(\p. \ (ofe_fields a) p \ \ (ofe_fields b) p))" unfolding check_no_overlap_alt check_no_overlap2_def by force text\Together with distinctness of the flow table, this provides the abscence of @{term Undefined}\footnote{It is slightly stronger than necessary, overlapping rules might be shadowed and thus never influence the behavior.}:\ lemma "\check_no_overlap \ ft; distinct ft\ \ OF_priority_match \ ft p \ Undefined" by (simp add: no_overlapsI no_overlaps_not_unefined) text\Given the absence of overlapping or duplicate flow entries, we can show two interesting equivalences. the first is the equality to the semantics defined by Guha \emph{et al.}:\ lemma "\check_no_overlap \ ft; distinct ft\ \ OF_priority_match \ ft p = option_to_ftb d \ guha_table_semantics \ ft p d" by (simp add: guha_equal no_overlapsI) text\where @{term option_to_ftb} maps between the return type of @{term OF_priority_match} and an option type as one would expect.\ text\The second equality for @{term OF_priority_match} is one that helps reasoning about flow tables. We define a simple recursive traversal for flow tables:\ lemma "OF_match_linear \ [] p = NoAction" "OF_match_linear \ (a#as) p = (if \ (ofe_fields a) p then Action (ofe_action a) else OF_match_linear \ as p)" by(fact OF_match_linear.simps)+ text\For this definition to be equivalent, we need the flow table to be sorted:\ lemma" \no_overlaps \ f ;sorted_descending (map ofe_prio f)\ \ OF_match_linear \ f p = OF_priority_match \ f p" by(fact OF_eq) text\As the last step, we implemented a serialization function for flow entries; it has to remain unverified. The serialization function deals with one little inaccuracy: We have modelled the @{term IngressPort} match to use the interface name, but OpenFlow requires numerical interface IDs instead. We deemed that pulling this translation step into the main translation would only make the correctness lemma of the translation more complicated while not increasing the confidence in the correctness significantly. We thus made replacing interface names by their ID part of the serialization. \ text\Having collected all important definitions and models, we can move on to the conversion.\ (*text\\todo{Maybe I should make a sweet little subsection that merges this all into a single model definition.}\*) subsection\Translation Implementation\ text_raw\\label{sec:convi}\ text\This section explains how the functions that are executed sequentially in a linux firewall can be compressed into a single OpenFlow table. Creating this flow table in a single step would be immensely complicated. We thus divided the task into several steps using the following key insights: \begin{itemize} \item All steps that are executed in the linux router can be formulated as a firewall, more specifically, a generalization of @{term simple_fw} that allows arbitrary actions instead of just accept and drop. \item A function that computes the conjunction of two @{term simple_fw} matches is already present. Extending this to a function that computes the join of two firewalls is relatively simple. This is explained in Section \ref{sec:fwconj} \end{itemize} \ subsubsection\Chaining Firewalls\ text_raw\\label{sec:fwconj}\ text\This section explains how to compute the join of two firewalls.\ text\The basis of this is a generalization of @{const simple_fw}. Instead of only allowing @{const simple_action.Accept} or @{const simple_action.Drop} as actions, it allows arbitrary actions. The type of the function that evaluates this generalized simple firewall is @{term "generalized_sfw :: ('i::len simple_match \ 'a) list \ ('i, 'b) simple_packet_scheme \ ('i simple_match \ 'a) option"}. The definition is straightforward:\ lemma "generalized_sfw [] p = None" "generalized_sfw (a # as) p = (if (case a of (m,_) \ simple_matches m p) then Some a else generalized_sfw as p)" by(fact generalized_sfw_simps)+ text\Based on that, we asked: if @{term fw\<^sub>1} makes the decision @{term a} (where @{term a} is the second element of the result tuple from @{const generalized_sfw}) and @{term fw\<^sub>2} makes the decision @{term b}, how can we compute the firewall that makes the decision @{term "(a,b)"}\footnote{Note that tuples are right-associative in Isabelle/HOL, i.e., @{term "(a::'a,(b,c)::('b\'c))"} is a pair of @{term a} and the pair @{term "(b,c)"}}. One possible answer is given by the following definition: \ lemma "generalized_fw_join l1 l2 \ [(u,a,b). (m1,a) \ l1, (m2,b) \ l2, u \ (case simple_match_and m1 m2 of None \ [] | Some s \ [s])]" by(fact generalized_fw_join_def[unfolded option2list_def])+ text\This definition validates the following lemma:\ lemma "generalized_sfw (generalized_fw_join fw\<^sub>1 fw\<^sub>2) p = Some (u, d\<^sub>1,d\<^sub>2) \ (\r\<^sub>1 r\<^sub>2. generalized_sfw fw\<^sub>1 p = Some (r\<^sub>1,d\<^sub>1) \ generalized_sfw fw\<^sub>2 p = Some (r\<^sub>2,d\<^sub>2) \ Some u = simple_match_and r\<^sub>1 r\<^sub>2)" by (auto dest: generalized_fw_joinD sym simp add: generalized_fw_joinI) text\Thus, @{const generalized_fw_join} has a number of applications. For example, it could be used to compute a firewall ruleset that represents two firewalls that are executed in sequence. \ definition "simple_action_conj a b \ (if a = simple_action.Accept \ b = simple_action.Accept then simple_action.Accept else simple_action.Drop)" definition "simple_rule_conj \ (uncurry SimpleRule \ apsnd (uncurry simple_action_conj))" theorem "simple_fw rs\<^sub>1 p = Decision FinalAllow \ simple_fw rs\<^sub>2 p = Decision FinalAllow \ simple_fw (map simple_rule_conj (generalized_fw_join (map simple_rule_dtor rs\<^sub>1) (map simple_rule_dtor rs\<^sub>2))) p = Decision FinalAllow" unfolding simple_rule_conj_def simple_action_conj_def[abs_def] using simple_fw_join by(force simp add: comp_def apsnd_def map_prod_def case_prod_unfold uncurry_def[abs_def]) text\Using the join, it should be possible to compute any $n$-ary logical operation on firewalls. We will use it for something somewhat different in the next section.\ subsubsection\Translation Implementation\ text_raw\ \begin{figure*} \begin{framed} \ lemma "lr_of_tran rt fw ifs \ if \ (no_oif_match fw \ has_default_policy fw \ simple_fw_valid fw \ valid_prefixes rt \ has_default_route rt \ distinct ifs) then Inl ''Error in creating OpenFlow table: prerequisites not satisifed'' else ( let nfw = map simple_rule_dtor fw; frt = map (\r. (route2match r, output_iface (routing_action r))) rt; nrd = generalized_fw_join frt nfw; ard = (map (apfst of_nat) \ annotate_rlen) nrd in - if length nrd < unat (max_word :: 16 word) + if length nrd < unat (- 1 :: 16 word) then Inr (pack_OF_entries ifs ard) else Inl ''Error in creating OpenFlow table: priority number space exhausted'' )" unfolding Let_def lr_of_tran_def lr_of_tran_fbs_def lr_of_tran_s1_def comp_def route2match_def by force text_raw\ \end{framed} \caption{Function for translating a @{typ "'i::len simple_rule list"}, a @{typ "'i routing_rule list"}, and a list of interfaces to a flow table.} \label{fig:convi} \end{figure*} \ text\ This section shows the actual definition of the translation function, in Figure~\ref{fig:convi}. Before beginning the translation, the definition checks whether the necessary preconditions are valid. This first two steps are to convert @{term fw} and @{term rt} to lists that can be evaluated by @{const generalized_sfw}. For @{term fw}, this is done by @{term "map simple_rule_dtor"}, which just deconstructs @{type simple_rule}s into tuples of match and action. For @{term rt}, we made a firewall ruleset with rules that use prefix matches on the destination IP address. The next step is to join the two rulesets. The result of the join is a ruleset with rules @{term r} that only match if both, the corresponding firewall rule @{term fwr} and the corresponding routing rule @{term rr} matches. The data accompanying @{term r} is the port from @{term rr} and the firewall decision from @{term fwr}. Next, descending priorities are added to the rules using @{term "map (apfst word_of_nat) \ annotate_rlen"}. If the number of rules is too large to fit into the $2^{16}$ priority classes, an error is returned. Otherwise, the function @{const pack_OF_entries} is used to convert the @{typ "(16 word \ 32 simple_match \ char list \ simple_action) list"} to an OpenFlow table. While converting the @{typ "char list \ simple_action"} tuple is straightforward, converting the @{type simple_match} to an equivalent list of @{typ "of_match_field set"} is non-trivial. This is done by the function @{const simple_match_to_of_match}. \ text\The main difficulties for @{const simple_match_to_of_match} lie in making sure that the prerequisites are satisfied and in the fact that a @{type simple_match} operates on slightly stronger match expressions. \begin{itemize} \item A @{type simple_match} allows a (string) prefix match on the input and output interfaces. Given a list of existing interfaces on the router @{term ifs}, the function has to insert flow entries for each interface matching the prefix. \item A @{type simple_match} can match ports by an interval. Now it becomes obvious why Section~\ref{sec:of_match} added bitmasks to @{const L4Src} and @{const L4Dst}. Using the algorithm to split word intervals into intervals that can be represented by prefix matches from~\cite{diekmann2016verified}, we can efficiently represent the original interval by a few (32 in the worst case) prefix matches and insert flow entries for each of them.% \footnote{It might be possible to represent the interval match more efficiently than a split into prefixes. However, that would produce overlapping matches (which is not a problem if we assing separate priorities) and we did not have a verified implementation of an algorithm that does so.} \end{itemize} The following lemma characterizes @{const simple_match_to_of_match}: \ lemma simple_match_to_of_match: assumes "simple_match_valid r" "p_iiface p \ set ifs" "match_iface (oiface r) (p_oiface p)" "p_l2type p = 0x800" shows "simple_matches r p \ (\gr \ set (simple_match_to_of_match r ifs). OF_match_fields gr p = Some True)" using assms simple_match_to_of_matchD simple_match_to_of_matchI by blast text\The assumptions are to be read as follows: \begin{itemize} \item The match @{term r} has to be valid, i.e. it has to use @{const valid_prefix} matches, and it cannot use anything other than $0$-$65535$ for the port matches unless its protocol match ensures @{const TCP}, @{const UDP} or @{const L4_Protocol.SCTP}. \item @{const simple_match_to_of_match} cannot produce rules for packets that have input interfaces that are not named in the interface list. \item The output interface of @{term p} has to match the output interface match of @{term r}. This is a weakened formulation of @{term "oiface r = ifaceAny"}, since @{thm[display] match_ifaceAny[no_vars]}. We require this because OpenFlow field matches cannot be used to match on the output port --- they are supposed to match a packet and decide an output port. \item The @{type simple_match} type was designed for IP(v4) packets, we limit ourselves to them. \end{itemize} The conclusion then states that the @{type simple_match} @{term r} matches iff an element of the result of @{const simple_match_to_of_match} matches. The third assumption is part of the explanation why we did not use @{const simple_linux_router_altered}: @{const simple_match_to_of_match} cannot deal with output interface matches. Thus, before passing a generalized simple firewall to @{const pack_OF_entries}, we would have to set the output ports to @{const ifaceAny}. A system replace output interface matches with destination IP addresses has already been formalized and will be published in a future version of \cite{Iptables_Semantics-AFP}. For now, we limit ourselves to firewalls that do not do output port matching, i.e., we require @{term "no_oif_match fw"}. \ text_raw\\begin{figure*} \begin{framed} \ theorem fixes p :: "(32, 'a) simple_packet_ext_scheme" assumes "p_iiface p \ set ifs" and "p_l2type p = 0x800" "lr_of_tran rt fw ifs = Inr oft" shows "OF_priority_match OF_match_fields_safe oft p = Action [Forward oif] \ simple_linux_router_nol12 rt fw p = (Some (p\p_oiface := oif\))" "OF_priority_match OF_match_fields_safe oft p = Action [] \ simple_linux_router_nol12 rt fw p = None" "OF_priority_match OF_match_fields_safe oft p \ NoAction" "OF_priority_match OF_match_fields_safe oft p \ Undefined" "OF_priority_match OF_match_fields_safe oft p = Action ls \ length ls \ 1" "\ls. length ls \ 1 \ OF_priority_match OF_match_fields_safe oft p = Action ls" using assms lr_of_tran_correct by simp_all text_raw\ \end{framed} \caption{Central theorem on @{const lr_of_tran}} \label{fig:central} \end{figure*} \ text\ Given discussed properties, we present the central theorem for our translation in Figure~\ref{fig:central}. The first two assumptions are limitations on the traffic we make a statement about. Obviously, we will never see any packets with an input interface that is not in the interface list. Furthermore, we do not state anything about non-IPv4 traffic. (The traffic will remain unmatched in by the flow table, but we have not verified that.) The last assumption is that the translation does not return a run-time error. The translation will return a run-time error if the rules can not be assigned priorities from a 16 bit integer, or when one of the following conditions on the input data is not satisifed:\ lemma " \ no_oif_match fw \ \ has_default_policy fw \ \ simple_fw_valid fw \ \ valid_prefixes rt \ \ has_default_route rt \ \ distinct ifs \ \err. lr_of_tran rt fw ifs = Inl err" unfolding lr_of_tran_def by(simp split: if_splits) subsubsection\Comparison to Exodus\ text\ We are not the first researchers to attempt automated static migration to SDN. The (only) other attempt we are aware of is \emph{Exodus} by Nelson \emph{et al.}~\cite{nelson2015exodus}. \ text\ There are some fundamental differences between Exodus and our work: \begin{itemize} \item Exodus focuses on Cisco IOS instead of linux. \item Exodus does not produce OpenFlow rulesets, but FlowLog~\cite{nelson2014tierless} controller programs. \item Exodus is not limited to using a single flow table. \item Exodus requires continuous controller interaction for some of its functions. \item Exodus attempts to support as much functionality as possible and has implemented support for dynamic routing, VLANs and NAT. \item Nelson \emph{et al.} reject the idea that the translation could or should be proven correct. \end{itemize} \ (*<*) end (*>*) diff --git a/thys/LOFT/OpenFlow_Serialize.thy b/thys/LOFT/OpenFlow_Serialize.thy --- a/thys/LOFT/OpenFlow_Serialize.thy +++ b/thys/LOFT/OpenFlow_Serialize.thy @@ -1,66 +1,66 @@ theory OpenFlow_Serialize imports OpenFlow_Matches OpenFlow_Action Semantics_OpenFlow Simple_Firewall.Primitives_toString IP_Addresses.Lib_Word_toString begin definition "serialization_test_entry \ OFEntry 7 {EtherDst 0x1, IPv4Dst (PrefixMatch 0xA000201 32), IngressPort ''s1-lan'', L4Dst 0x50 0, L4Src 0x400 0x3FF, IPv4Proto 6, EtherType 0x800} [ModifyField_l2dst 0xA641F185E862, Forward ''s1-wan'']" value "(map ((<<) (1::48 word) \ (*) 8) \ rev) [0..<6]" definition "serialize_mac (m::48 word) \ (intersperse (CHR '':'') \ map (hex_string_of_word 1 \ (\h. (m >> h * 8) && 0xff)) \ rev) [0..<6]" lemma "serialize_mac 0xdeadbeefcafe = ''de:ad:be:ef:ca:fe''" by eval definition "serialize_action pids a \ (case a of Forward oif \ ''output:'' @ pids oif | ModifyField_l2dst na \ ''mod_dl_dst:'' @ serialize_mac na)" definition "serialize_actions pids a \ if length a = 0 then ''drop'' else (intersperse (CHR '','') \ map (serialize_action pids)) a" lemma "serialize_actions (\oif. ''42'') (ofe_action serialization_test_entry) = ''mod_dl_dst:a6:41:f1:85:e8:62,output:42''" by eval lemma "serialize_actions anything [] = ''drop''" by(simp add: serialize_actions_def) definition "prefix_to_string pfx \ ipv4_cidr_toString (pfxm_prefix pfx, pfxm_length pfx)" primrec serialize_of_match where "serialize_of_match pids (IngressPort p) = ''in_port='' @ pids p" | "serialize_of_match _ (VlanId i) = ''dl_vlan='' @ dec_string_of_word0 i" | "serialize_of_match _ (VlanPriority _) = undefined" | (* uh, äh\ We don't use that anyway\ *) "serialize_of_match _ (EtherType i) = ''dl_type=0x'' @ hex_string_of_word0 i" | "serialize_of_match _ (EtherSrc m) = ''dl_src='' @ serialize_mac m" | "serialize_of_match _ (EtherDst m) = ''dl_dst='' @ serialize_mac m" | "serialize_of_match _ (IPv4Proto i) = ''nw_proto='' @ dec_string_of_word0 i" | "serialize_of_match _ (IPv4Src p) = ''nw_src='' @ prefix_to_string p" | "serialize_of_match _ (IPv4Dst p) = ''nw_dst='' @ prefix_to_string p" | -"serialize_of_match _ (L4Src i m) = ''tp_src='' @ dec_string_of_word0 i @ (if m = max_word then [] else ''/0x'' @ hex_string_of_word 3 m)" | -"serialize_of_match _ (L4Dst i m) = ''tp_dst='' @ dec_string_of_word0 i @ (if m = max_word then [] else ''/0x'' @ hex_string_of_word 3 m)" +"serialize_of_match _ (L4Src i m) = ''tp_src='' @ dec_string_of_word0 i @ (if m = - 1 then [] else ''/0x'' @ hex_string_of_word 3 m)" | +"serialize_of_match _ (L4Dst i m) = ''tp_dst='' @ dec_string_of_word0 i @ (if m = - 1 then [] else ''/0x'' @ hex_string_of_word 3 m)" definition serialize_of_matches :: "(string \ string) \ of_match_field set \ string" where "serialize_of_matches pids \ (@) ''hard_timeout=0,idle_timeout=0,'' \ intersperse (CHR '','') \ map (serialize_of_match pids) \ sorted_list_of_set" lemma "serialize_of_matches pids of_matches= (List.append ''hard_timeout=0,idle_timeout=0,'') (intersperse (CHR '','') (map (serialize_of_match pids) (sorted_list_of_set of_matches)))" by (simp add: serialize_of_matches_def) export_code serialize_of_matches checking SML (*needs "HOL-Library.Code_Char"*) lemma "serialize_of_matches (\oif. ''42'') (ofe_fields serialization_test_entry) = ''hard_timeout=0,idle_timeout=0,in_port=42,dl_type=0x800,dl_dst=00:00:00:00:00:01,nw_proto=6,nw_dst=10.0.2.1/32,tp_src=1024/0x03ff,tp_dst=80/0x0000''" by eval definition "serialize_of_entry pids e \ (case e of (OFEntry p f a) \ ''priority='' @ dec_string_of_word0 p @ '','' @ serialize_of_matches pids f @ '','' @ ''action='' @ serialize_actions pids a)" lemma "serialize_of_entry (the \ map_of [(''s1-lan'',''42''),(''s1-wan'',''1337'')]) serialization_test_entry = ''priority=7,hard_timeout=0,idle_timeout=0,in_port=42,dl_type=0x800,dl_dst=00:00:00:00:00:01,nw_proto=6,nw_dst=10.0.2.1/32,tp_src=1024/0x03ff,tp_dst=80/0x0000,action=mod_dl_dst:a6:41:f1:85:e8:62,output:1337''" by eval end diff --git a/thys/MFODL_Monitor_Optimized/Code_Double.thy b/thys/MFODL_Monitor_Optimized/Code_Double.thy --- a/thys/MFODL_Monitor_Optimized/Code_Double.thy +++ b/thys/MFODL_Monitor_Optimized/Code_Double.thy @@ -1,521 +1,521 @@ (*<*) theory Code_Double imports IEEE_Floating_Point.IEEE_Properties "HOL-Library.Code_Target_Int" Containers.Collection_Eq Containers.Collection_Order begin (*>*) section \Code adaptation for IEEE double-precision floats\ subsection \copysign\ lift_definition copysign :: "('e, 'f) float \ ('e, 'f) float \ ('e, 'f) float" is "\(_, e::'e word, f::'f word) (s::1 word, _, _). (s, e, f)" . lemma is_nan_copysign[simp]: "is_nan (copysign x y) \ is_nan x" unfolding is_nan_def by transfer auto subsection \Additional lemmas about generic floats\ lemma is_nan_some_nan[simp]: "is_nan (some_nan :: ('e, 'f) float)" unfolding some_nan_def - by (rule someI[where x="Abs_float (0, max_word :: 'e word, 1)"]) + by (rule someI[where x="Abs_float (0, - 1 :: 'e word, 1)"]) (auto simp add: is_nan_def exponent_def fraction_def emax_def Abs_float_inverse[simplified]) lemma not_is_nan_0[simp]: "\ is_nan 0" unfolding is_nan_def by (simp add: zero_simps) lemma not_is_nan_1[simp]: "\ is_nan 1" unfolding is_nan_def by transfer simp lemma is_nan_plus: "is_nan x \ is_nan y \ is_nan (x + y)" unfolding plus_float_def fadd_def by auto lemma is_nan_minus: "is_nan x \ is_nan y \ is_nan (x - y)" unfolding minus_float_def fsub_def by auto lemma is_nan_times: "is_nan x \ is_nan y \ is_nan (x * y)" unfolding times_float_def fmul_def by auto lemma is_nan_divide: "is_nan x \ is_nan y \ is_nan (x / y)" unfolding divide_float_def fdiv_def by auto lemma is_nan_float_sqrt: "is_nan x \ is_nan (float_sqrt x)" unfolding float_sqrt_def fsqrt_def by simp lemma nan_fcompare: "is_nan x \ is_nan y \ fcompare x y = Und" unfolding fcompare_def by simp lemma nan_not_le: "is_nan x \ is_nan y \ \ x \ y" unfolding less_eq_float_def fle_def fcompare_def by simp lemma nan_not_less: "is_nan x \ is_nan y \ \ x < y" unfolding less_float_def flt_def fcompare_def by simp lemma nan_not_zero: "is_nan x \ \ is_zero x" unfolding is_nan_def is_zero_def by simp lemma nan_not_infinity: "is_nan x \ \ is_infinity x" unfolding is_nan_def is_infinity_def by simp lemma zero_not_infinity: "is_zero x \ \ is_infinity x" unfolding is_zero_def is_infinity_def by simp lemma zero_not_nan: "is_zero x \ \ is_nan x" unfolding is_zero_def is_nan_def by simp lemma minus_one_power_one_word: "(-1 :: real) ^ unat (x :: 1 word) = (if unat x = 0 then 1 else -1)" proof - have "unat x = 0 \ unat x = 1" using le_SucE[OF unat_one_word_le] by auto then show ?thesis by auto qed definition valofn :: "('e, 'f) float \ real" where "valofn x = (2^exponent x / 2^bias TYPE(('e, 'f) float)) * (1 + real (fraction x) / 2^LENGTH('f))" definition valofd :: "('e, 'f) float \ real" where "valofd x = (2 / 2^bias TYPE(('e, 'f) float)) * (real (fraction x) / 2^LENGTH('f))" lemma valof_alt: "valof x = (if exponent x = 0 then if sign x = 0 then valofd x else - valofd x else if sign x = 0 then valofn x else - valofn x)" unfolding valofn_def valofd_def by transfer (auto simp: minus_one_power_one_word unat_eq_0 field_simps) lemma fraction_less_2p: "fraction (x :: ('e, 'f) float) < 2^LENGTH('f)" by transfer auto lemma valofn_ge_0: "0 \ valofn x" unfolding valofn_def by simp lemma valofn_ge_2p: "2^exponent (x :: ('e, 'f) float) / 2^bias TYPE(('e, 'f) float) \ valofn x" unfolding valofn_def by (simp add: field_simps) lemma valofn_less_2p: fixes x :: "('e, 'f) float" assumes "exponent x < e" shows "valofn x < 2^e / 2^bias TYPE(('e, 'f) float)" proof - have "1 + real (fraction x) / 2^LENGTH('f) < 2" by (simp add: fraction_less_2p) then have "valofn x < (2^exponent x / 2^bias TYPE(('e, 'f) float)) * 2" unfolding valofn_def by (simp add: field_simps) also have "... \ 2^e / 2^bias TYPE(('e, 'f) float)" using assms by (auto simp: less_eq_Suc_le field_simps elim: order_trans[rotated, OF exp_less]) finally show ?thesis . qed lemma valofd_ge_0: "0 \ valofd x" unfolding valofd_def by simp lemma valofd_less_2p: "valofd (x :: ('e, 'f) float) < 2 / 2^bias TYPE(('e, 'f) float)" unfolding valofd_def by (simp add: fraction_less_2p field_simps) lemma valofn_le_imp_exponent_le: fixes x y :: "('e, 'f) float" assumes "valofn x \ valofn y" shows "exponent x \ exponent y" proof (rule ccontr) assume "\ exponent x \ exponent y" then have "valofn y < 2^exponent x / 2^bias TYPE(('e, 'f) float)" using valofn_less_2p[of y] by auto also have "... \ valofn x" by (rule valofn_ge_2p) finally show False using assms by simp qed lemma valofn_eq: fixes x y :: "('e, 'f) float" assumes "valofn x = valofn y" shows "exponent x = exponent y" "fraction x = fraction y" proof - from assms show "exponent x = exponent y" by (auto intro!: antisym valofn_le_imp_exponent_le) with assms show "fraction x = fraction y" unfolding valofn_def by (simp add: field_simps) qed lemma valofd_eq: fixes x y :: "('e, 'f) float" assumes "valofd x = valofd y" shows "fraction x = fraction y" using assms unfolding valofd_def by (simp add: field_simps) lemma is_zero_valof_conv: "is_zero x \ valof x = 0" unfolding is_zero_def valof_alt using valofn_ge_2p[of x] by (auto simp: valofd_def field_simps dest: order.antisym) lemma valofd_neq_valofn: fixes x y :: "('e, 'f) float" assumes "exponent y \ 0" shows "valofd x \ valofn y" "valofn y \ valofd x" proof - have "valofd x < 2 / 2^bias TYPE(('e, 'f) float)" by (rule valofd_less_2p) also have "... \ 2 ^ IEEE.exponent y / 2 ^ bias TYPE(('e, 'f) IEEE.float)" using assms by (simp add: self_le_power field_simps) also have "... \ valofn y" by (rule valofn_ge_2p) finally show "valofd x \ valofn y" "valofn y \ valofd x" by simp_all qed lemma sign_gt_0_conv: "0 < sign x \ sign x = 1" by (cases x rule: sign_cases) simp_all lemma valof_eq: assumes "\ is_zero x \ \ is_zero y" shows "valof x = valof y \ x = y" proof assume *: "valof x = valof y" with assms have "valof y \ 0" by (simp add: is_zero_valof_conv) with * show "x = y" unfolding valof_alt using valofd_ge_0[of x] valofd_ge_0[of y] valofn_ge_0[of x] valofn_ge_0[of y] by (auto simp: valofd_neq_valofn sign_gt_0_conv split: if_splits intro!: float_eqI elim: valofn_eq valofd_eq) qed simp lemma zero_fcompare: "is_zero x \ is_zero y \ fcompare x y = ccode.Eq" unfolding fcompare_def by (simp add: zero_not_infinity zero_not_nan is_zero_valof_conv) subsection \Doubles with a unified NaN value\ quotient_type double = "(11, 52) float" / "\x y. is_nan x \ is_nan y \ x = y" by (auto intro!: equivpI reflpI sympI transpI) instantiation double :: "{zero, one, plus, minus, uminus, times, ord}" begin lift_definition zero_double :: "double" is "0" . lift_definition one_double :: "double" is "1" . lift_definition plus_double :: "double \ double \ double" is plus by (auto simp: is_nan_plus) lift_definition minus_double :: "double \ double \ double" is minus by (auto simp: is_nan_minus) lift_definition uminus_double :: "double \ double" is uminus by auto lift_definition times_double :: "double \ double \ double" is times by (auto simp: is_nan_times) lift_definition less_eq_double :: "double \ double \ bool" is "(\)" by (auto simp: nan_not_le) lift_definition less_double :: "double \ double \ bool" is "(<)" by (auto simp: nan_not_less) instance .. end instantiation double :: inverse begin lift_definition divide_double :: "double \ double \ double" is divide by (auto simp: is_nan_divide) definition inverse_double :: "double \ double" where "inverse_double x = 1 div x" instance .. end lift_definition sqrt_double :: "double \ double" is float_sqrt by (auto simp: is_nan_float_sqrt) no_notation plus_infinity ("\") lift_definition infinity :: "double" is plus_infinity . lift_definition nan :: "double" is some_nan . lift_definition is_zero :: "double \ bool" is IEEE.is_zero by (auto simp: nan_not_zero) lift_definition is_infinite :: "double \ bool" is IEEE.is_infinity by (auto simp: nan_not_infinity) lift_definition is_nan :: "double \ bool" is IEEE.is_nan by auto lemma is_nan_conv: "is_nan x \ x = nan" by transfer auto lift_definition copysign_double :: "double \ double \ double" is "\x y. if IEEE.is_nan y then some_nan else copysign x y" by auto text \Note: @{const copysign_double} deviates from the IEEE standard in cases where the second argument is a NaN.\ lift_definition fcompare_double :: "double \ double \ ccode" is fcompare by (auto simp: nan_fcompare) lemma nan_fcompare_double: "is_nan x \ is_nan y \ fcompare_double x y = Und" by transfer (rule nan_fcompare) consts compare_double :: "double \ double \ integer" specification (compare_double) compare_double_less: "compare_double x y < 0 \ is_nan x \ \ is_nan y \ fcompare_double x y = ccode.Lt" compare_double_eq: "compare_double x y = 0 \ is_nan x \ is_nan y \ fcompare_double x y = ccode.Eq" compare_double_greater: "compare_double x y > 0 \ \ is_nan x \ is_nan y \ fcompare_double x y = ccode.Gt" by (rule exI[where x="\x y. if is_nan x then if is_nan y then 0 else -1 else if is_nan y then 1 else (case fcompare_double x y of ccode.Eq \ 0 | ccode.Lt \ -1 | ccode.Gt \ 1)"], transfer, auto simp: fcompare_def) lemmas compare_double_simps = compare_double_less compare_double_eq compare_double_greater lemma compare_double_le_0: "compare_double x y \ 0 \ is_nan x \ fcompare_double x y \ {ccode.Eq, ccode.Lt}" by (rule linorder_cases[of "compare_double x y" 0]; simp) (auto simp: compare_double_simps nan_fcompare_double) lift_definition double_of_integer :: "integer \ double" is "\x. zerosign 0 (intround To_nearest (int_of_integer x))" . definition double_of_int where [code del]: "double_of_int x = double_of_integer (integer_of_int x)" lemma [code]: "double_of_int (int_of_integer x) = double_of_integer x" unfolding double_of_int_def by simp lift_definition integer_of_double :: "double \ integer" is "\x. if IEEE.is_nan x \ IEEE.is_infinity x then undefined else integer_of_int \valof (intround float_To_zero (valof x) :: (11, 52) float)\" by auto definition int_of_double: "int_of_double x = int_of_integer (integer_of_double x)" subsection \Linear ordering\ definition lcompare_double :: "double \ double \ integer" where "lcompare_double x y = (if is_zero x \ is_zero y then compare_double (copysign_double 1 x) (copysign_double 1 y) else compare_double x y)" lemma fcompare_double_swap: "fcompare_double x y = ccode.Gt \ fcompare_double y x = ccode.Lt" by transfer (auto simp: fcompare_def) lemma fcompare_double_refl: "\ is_nan x \ fcompare_double x x = ccode.Eq" by transfer (auto simp: fcompare_def) lemma fcompare_double_Eq1: "fcompare_double x y = ccode.Eq \ fcompare_double y z = c \ fcompare_double x z = c" by transfer (auto simp: fcompare_def split: if_splits) lemma fcompare_double_Eq2: "fcompare_double y z = ccode.Eq \ fcompare_double x y = c \ fcompare_double x z = c" by transfer (auto simp: fcompare_def split: if_splits) lemma fcompare_double_Lt_trans: "fcompare_double x y = ccode.Lt \ fcompare_double y z = ccode.Lt \ fcompare_double x z = ccode.Lt" by transfer (auto simp: fcompare_def split: if_splits) lemma fcompare_double_eq: "\ is_zero x \ \ is_zero y \ fcompare_double x y = ccode.Eq \ x = y" by transfer (auto simp: fcompare_def valof_eq IEEE.is_infinity_def split: if_splits intro!: float_eqI) lemma fcompare_double_Lt_asym: "fcompare_double x y = ccode.Lt \ fcompare_double y x = ccode.Lt \ False" by transfer (auto simp: fcompare_def split: if_splits) lemma compare_double_swap: "0 < compare_double x y \ compare_double y x < 0" by (auto simp: compare_double_simps fcompare_double_swap) lemma compare_double_refl: "compare_double x x = 0" by (auto simp: compare_double_eq intro!: fcompare_double_refl) lemma compare_double_trans: "compare_double x y \ 0 \ compare_double y z \ 0 \ compare_double x z \ 0" by (fastforce simp: compare_double_le_0 nan_fcompare_double dest: fcompare_double_Eq1 fcompare_double_Eq2 fcompare_double_Lt_trans) lemma compare_double_antisym: "compare_double x y \ 0 \ compare_double y x \ 0 \ \ is_zero x \ \ is_zero y \ x = y" unfolding compare_double_le_0 by (auto simp: nan_fcompare_double is_nan_conv intro: fcompare_double_eq fcompare_double_eq[symmetric] dest: fcompare_double_Lt_asym) lemma zero_compare_double_copysign: "compare_double (copysign_double 1 x) (copysign_double 1 y) \ 0 \ is_zero x \ is_zero y \ compare_double x y \ 0" unfolding compare_double_le_0 by transfer (auto simp: nan_not_zero zero_fcompare split: if_splits) lemma is_zero_double_cases: "is_zero x \ (x = 0 \ P) \ (x = -0 \ P) \ P" by transfer (auto elim!: is_zero_cases) lemma copysign_1_0[simp]: "copysign_double 1 0 = 1" "copysign_double 1 (-0) = -1" by (transfer, simp, transfer, auto)+ lemma is_zero_uminus_double[simp]: "is_zero (- x) \ is_zero x" by transfer simp lemma not_is_zero_one_double[simp]: "\ is_zero 1" by (transfer, unfold IEEE.is_zero_def, transfer, simp) lemma uminus_one_neq_one_double[simp]: "- 1 \ (1 :: double)" by (transfer, transfer, simp) definition lle_double :: "double \ double \ bool" where "lle_double x y \ lcompare_double x y \ 0" definition lless_double :: "double \ double \ bool" where "lless_double x y \ lcompare_double x y < 0" lemma lcompare_double_ge_0: "lcompare_double x y \ 0 \ lle_double y x" unfolding lle_double_def lcompare_double_def using compare_double_swap not_less by auto lemma lcompare_double_gt_0: "lcompare_double x y > 0 \ lless_double y x" unfolding lless_double_def lcompare_double_def using compare_double_swap by auto lemma lcompare_double_eq_0: "lcompare_double x y = 0 \ x = y" proof assume *: "lcompare_double x y = 0" show "x = y" proof (cases "is_zero x \ is_zero y") case True with * show ?thesis by (fastforce simp: lcompare_double_def compare_double_simps is_nan_conv elim: is_zero_double_cases dest!: fcompare_double_eq[rotated]) next case False with * show ?thesis by (auto simp: lcompare_double_def linorder_not_less[symmetric] compare_double_swap intro!: compare_double_antisym) qed next assume "x = y" then show "lcompare_double x y = 0" by (simp add: lcompare_double_def compare_double_refl) qed lemmas lcompare_double_0_folds = lle_double_def[symmetric] lless_double_def[symmetric] lcompare_double_ge_0 lcompare_double_gt_0 lcompare_double_eq_0 interpretation double_linorder: linorder lle_double lless_double proof fix x y z :: double show "lless_double x y \ lle_double x y \ \ lle_double y x" unfolding lless_double_def lle_double_def lcompare_double_def by (auto simp: compare_double_swap not_le) show "lle_double x x" unfolding lle_double_def lcompare_double_def by (auto simp: compare_double_refl) show "lle_double x z" if "lle_double x y" and "lle_double y z" using that by (auto 0 3 simp: lle_double_def lcompare_double_def not_le compare_double_swap split: if_splits dest: compare_double_trans zero_compare_double_copysign zero_compare_double_copysign[OF less_imp_le] compare_double_antisym) show "x = y" if "lle_double x y" and "lle_double y x" proof (cases "is_zero x \ is_zero y") case True with that show ?thesis by (auto 0 3 simp: lle_double_def lcompare_double_def elim: is_zero_double_cases dest!: compare_double_antisym) next case False with that show ?thesis by (auto simp: lle_double_def lcompare_double_def elim!: compare_double_antisym) qed show "lle_double x y \ lle_double y x" unfolding lle_double_def lcompare_double_def by (auto simp: compare_double_swap not_le) qed instantiation double :: equal begin definition equal_double :: "double \ double \ bool" where "equal_double x y \ lcompare_double x y = 0" instance by intro_classes (simp add: equal_double_def lcompare_double_eq_0) end derive (eq) ceq double definition comparator_double :: "double comparator" where "comparator_double x y = (let c = lcompare_double x y in if c = 0 then order.Eq else if c < 0 then order.Lt else order.Gt)" lemma comparator_double: "comparator comparator_double" unfolding comparator_double_def by (auto simp: lcompare_double_0_folds split: if_splits intro!: comparator.intro) local_setup \ Comparator_Generator.register_foreign_comparator @{typ double} @{term comparator_double} @{thm comparator_double} \ derive ccompare double subsubsection \Code setup\ declare [[code drop: "0 :: double" "1 :: double" "plus :: double \ _" "minus :: double \ _" "uminus :: double \ _" "times :: double \ _" "less_eq :: double \ _" "less :: double \ _" "divide :: double \ _" sqrt_double infinity nan is_zero is_infinite is_nan copysign_double fcompare_double double_of_integer integer_of_double ]] code_printing code_module FloatUtil \ (OCaml) \module FloatUtil : sig val iszero : float -> bool val isinfinite : float -> bool val isnan : float -> bool val copysign : float -> float -> float val compare : float -> float -> Z.t end = struct let iszero x = (Pervasives.classify_float x = Pervasives.FP_zero);; let isinfinite x = (Pervasives.classify_float x = Pervasives.FP_infinite);; let isnan x = (Pervasives.classify_float x = Pervasives.FP_nan);; let copysign x y = if isnan y then Pervasives.nan else Pervasives.copysign x y;; let compare x y = Z.of_int (Pervasives.compare x y);; end;;\ code_reserved OCaml Pervasives FloatUtil code_printing type_constructor double \ (OCaml) "float" | constant "uminus :: double \ double" \ (OCaml) "Pervasives.(~-.)" | constant "(+) :: double \ double \ double" \ (OCaml) "Pervasives.(+.)" | constant "(*) :: double \ double \ double" \ (OCaml) "Pervasives.( *. )" | constant "(/) :: double \ double \ double" \ (OCaml) "Pervasives.('/.)" | constant "(-) :: double \ double \ double" \ (OCaml) "Pervasives.(-.)" | constant "0 :: double" \ (OCaml) "0.0" | constant "1 :: double" \ (OCaml) "1.0" | constant "(\) :: double \ double \ bool" \ (OCaml) "Pervasives.(<=)" | constant "(<) :: double \ double \ bool" \ (OCaml) "Pervasives.(<)" | constant "sqrt_double :: double \ double" \ (OCaml) "Pervasives.sqrt" | constant "infinity :: double" \ (OCaml) "Pervasives.infinity" | constant "nan :: double" \ (OCaml) "Pervasives.nan" | constant "is_zero :: double \ bool" \ (OCaml) "FloatUtil.iszero" | constant "is_infinite :: double \ bool" \ (OCaml) "FloatUtil.isinfinite" | constant "is_nan :: double \ bool" \ (OCaml) "FloatUtil.isnan" | constant "copysign_double :: double \ double \ double" \ (OCaml) "FloatUtil.copysign" | constant "compare_double :: double \ double \ integer" \ (OCaml) "FloatUtil.compare" | constant "double_of_integer :: integer \ double" \ (OCaml) "Z.to'_float" | constant "integer_of_double :: double \ integer" \ (OCaml) "Z.of'_float" hide_const (open) fcompare_double (*<*) end (*>*) diff --git a/thys/Simple_Firewall/Primitives/Primitives_toString.thy b/thys/Simple_Firewall/Primitives/Primitives_toString.thy --- a/thys/Simple_Firewall/Primitives/Primitives_toString.thy +++ b/thys/Simple_Firewall/Primitives/Primitives_toString.thy @@ -1,58 +1,58 @@ section\toString Functions for Primitives\ theory Primitives_toString imports "../Common/Lib_Enum_toString" IP_Addresses.IP_Address_toString Iface L4_Protocol begin definition ipv4_cidr_toString :: "(ipv4addr \ nat) \ string" where "ipv4_cidr_toString ip_n = (case ip_n of (base, n) \ (ipv4addr_toString base @''/''@ string_of_nat n))" lemma "ipv4_cidr_toString (ipv4addr_of_dotdecimal (192,168,0,1), 22) = ''192.168.0.1/22''" by eval definition ipv6_cidr_toString :: "(ipv6addr \ nat) \ string" where "ipv6_cidr_toString ip_n = (case ip_n of (base, n) \ (ipv6addr_toString base @''/''@ string_of_nat n))" lemma "ipv6_cidr_toString (42540766411282592856906245548098208122, 64) = ''2001:db8::8:800:200c:417a/64''" by eval definition primitive_protocol_toString :: "primitive_protocol \ string" where "primitive_protocol_toString protid \ ( if protid = TCP then ''tcp'' else if protid = UDP then ''udp'' else if protid = ICMP then ''icmp'' else if protid = L4_Protocol.SCTP then ''sctp'' else if protid = L4_Protocol.IGMP then ''igmp'' else if protid = L4_Protocol.GRE then ''gre'' else if protid = L4_Protocol.ESP then ''esp'' else if protid = L4_Protocol.AH then ''ah'' else if protid = L4_Protocol.IPv6ICMP then ''ipv6-icmp'' else ''protocolid:''@dec_string_of_word0 protid)" fun protocol_toString :: "protocol \ string" where "protocol_toString (ProtoAny) = ''all''" | "protocol_toString (Proto protid) = primitive_protocol_toString protid" definition iface_toString :: "string \ iface \ string" where "iface_toString descr iface = (if iface = ifaceAny then '''' else (case iface of (Iface name) \ descr@name))" lemma "iface_toString ''in: '' (Iface ''+'') = ''''" by eval lemma "iface_toString ''in: '' (Iface ''eth0'') = ''in: eth0''" by eval definition port_toString :: "16 word \ string" where "port_toString p \ dec_string_of_word0 p" fun ports_toString :: "string \ (16 word \ 16 word) \ string" where - "ports_toString descr (s,e) = (if s = 0 \ e = max_word then '''' else descr @ (if s=e then port_toString s else port_toString s@'':''@port_toString e))" + "ports_toString descr (s,e) = (if s = 0 \ e = - 1 then '''' else descr @ (if s=e then port_toString s else port_toString s@'':''@port_toString e))" lemma "ports_toString ''spt: '' (0,65535) = ''''" by eval lemma "ports_toString ''spt: '' (1024,2048) = ''spt: 1024:2048''" by eval lemma "ports_toString ''spt: '' (1024,1024) = ''spt: 1024''" by eval definition ipv4_cidr_opt_toString :: "string \ ipv4addr \ nat \ string" where "ipv4_cidr_opt_toString descr ip = (if ip = (0,0) then '''' else descr@ipv4_cidr_toString ip)" definition protocol_opt_toString :: "string \ protocol \ string" where "protocol_opt_toString descr prot = (if prot = ProtoAny then '''' else descr@protocol_toString prot)" end diff --git a/thys/Simple_Firewall/Service_Matrix.thy b/thys/Simple_Firewall/Service_Matrix.thy --- a/thys/Simple_Firewall/Service_Matrix.thy +++ b/thys/Simple_Firewall/Service_Matrix.thy @@ -1,1661 +1,1661 @@ (* Title: Service_Matrix.thy Authors: Cornelius Diekmann, Max Haslbeck *) (*IPPartitioning.thy Original Author: Max Haslbeck, 2015*) section\Service Matrices\ theory Service_Matrix imports "Common/List_Product_More" "Common/IP_Partition_Preliminaries" "Common/GroupF" "Common/IP_Addr_WordInterval_toString" "Primitives/Primitives_toString" SimpleFw_Semantics IP_Addresses.WordInterval_Sorted begin subsection\IP Address Space Partition\ (* could be generalized more *) fun extract_IPSets_generic0 :: "('i::len simple_match \ 'i word \ nat) \ 'i simple_rule list \ ('i wordinterval) list" where "extract_IPSets_generic0 _ [] = []" | "extract_IPSets_generic0 sel ((SimpleRule m _)#ss) = (ipcidr_tuple_to_wordinterval (sel m)) # (extract_IPSets_generic0 sel ss)" lemma extract_IPSets_generic0_length: "length (extract_IPSets_generic0 sel rs) = length rs" by(induction rs rule: extract_IPSets_generic0.induct) (simp_all) (* The order in which extract_IPSets returns the collected IP ranges heavily influences the running time of the subsequent algorithms. For example: 1) iterating through the ruleset and collecting all source and destination ips: 10:10:49 elapsed time, 38:41:17 cpu time, factor 3.80 2) iterating through the ruleset and first returning all source ips and iterating again and then return all destination ips: 3:39:56 elapsed time, 21:08:34 cpu time, factor 5.76 To get a more deterministic runtime, we are sorting the output. As a performance optimization, we also remove duplicate entries. We use mergesort_remdups, which does a mergesort (i.e sorts!) and removes duplicates and mergesort_by_rel which does a mergesort (without removing duplicates) and allows to specify the relation we use to sort. In theory, the largest ip ranges (smallest prefix length) should be put first, the following evaluation shows that this may not be the fastest solution. The reason might be that access_matrix_pretty_ipv4 picks (almost randomly) one IP from the result and there are fast and slower choices. The faster choices are the ones where the firewall ruleset has a decision very early. Therefore, the running time is still a bit unpredictable. Here is the data: map ipcidr_tuple_to_wordinterval (mergesort_by_rel (\ (a1,a2) (b1, b2). (a2, a1) \ (b2, b1)) (mergesort_remdups ((map (src \ match_sel) rs) @ (map (dst \ match_sel) rs)))) (2:47:04 elapsed time, 17:08:01 cpu time, factor 6.15) map ipcidr_tuple_to_wordinterval (mergesort_remdups ((map (src \ match_sel) rs) @ (map (dst \ match_sel) rs))) (2:41:03 elapsed time, 16:56:46 cpu time, factor 6.31) map ipcidr_tuple_to_wordinterval (mergesort_by_rel (\ (a1,a2) (b1, b2). (a2, a1) \ (b2, b1)) ( ((map (src \ match_sel) rs) @ (map (dst \ match_sel) rs))) (5:52:28 elapsed time, 41:50:10 cpu time, factor 7.12) map ipcidr_tuple_to_wordinterval (mergesort_by_rel (\) ((map (src \ match_sel) rs) @ (map (dst \ match_sel) rs)))) (3:10:57 elapsed time, 19:12:25 cpu time, factor 6.03) map ipcidr_tuple_to_wordinterval (mergesort_by_rel (\ (a1,a2) (b1, b2). (a2, a1) \ (b2, b1)) (mergesort_remdups (extract_src_dst_ips rs []))) (2:49:57 elapsed time, 17:10:49 cpu time, factor 6.06) map ipcidr_tuple_to_wordinterval ((mergesort_remdups (extract_src_dst_ips rs []))) (2:43:44 elapsed time, 16:57:49 cpu time, factor 6.21) map ipcidr_tuple_to_wordinterval (mergesort_by_rel (\ (a1,a2) (b1, b2). (a2, a1) \ (b2, b1)) (mergesort_remdups (extract_src_dst_ips rs []))) (2:47:37 elapsed time, 16:54:47 cpu time, factor 6.05) There is a clear looser: not using mergesort_remdups There is no clear winner. We will just stick to mergesort_remdups. *) (*check the the order of mergesort_remdups did not change*) lemma "mergesort_remdups [(1::ipv4addr, 2::nat), (8,0), (8,1), (2,2), (2,4), (1,2), (2,2)] = [(1, 2), (2, 2), (2, 4), (8, 0), (8, 1)]" by eval (*a tail-recursive implementation*) fun extract_src_dst_ips :: "'i::len simple_rule list \ ('i word \ nat) list \ ('i word \ nat) list" where "extract_src_dst_ips [] ts = ts" | "extract_src_dst_ips ((SimpleRule m _)#ss) ts = extract_src_dst_ips ss (src m # dst m # ts)" lemma extract_src_dst_ips_length: "length (extract_src_dst_ips rs acc) = 2*length rs + length acc" proof(induction rs arbitrary: acc) case (Cons r rs) thus ?case by(cases r, simp) qed(simp) definition extract_IPSets :: "'i::len simple_rule list \ ('i wordinterval) list" where "extract_IPSets rs \ map ipcidr_tuple_to_wordinterval (mergesort_remdups (extract_src_dst_ips rs []))" lemma extract_IPSets: "set (extract_IPSets rs) = set (extract_IPSets_generic0 src rs) \ set (extract_IPSets_generic0 dst rs)" proof - { fix acc have "ipcidr_tuple_to_wordinterval ` set (extract_src_dst_ips rs acc) = ipcidr_tuple_to_wordinterval ` set acc \ set (extract_IPSets_generic0 src rs) \ set (extract_IPSets_generic0 dst rs)" proof(induction rs arbitrary: acc) case (Cons r rs ) thus ?case apply(cases r) apply(simp) by fast qed(simp) } thus ?thesis unfolding extract_IPSets_def by(simp_all add: extract_IPSets_def mergesort_remdups_correct) qed lemma "(a::nat) div 2 + a mod 2 \ a" by fastforce lemma merge_length: "length (merge l1 l2) \ length l1 + length l2" by(induction l1 l2 rule: merge.induct) auto lemma merge_list_length: "length (merge_list as ls) \ length (concat (as @ ls))" proof(induction as ls rule: merge_list.induct) case (5 l1 l2 acc2 ls) have "length (merge l2 acc2) \ length l2 + length acc2" using merge_length by blast with 5 show ?case by simp qed(simp_all) lemma mergesort_remdups_length: "length (mergesort_remdups as) \ length as" unfolding mergesort_remdups_def proof - have "concat ([] @ (map (\x. [x]) as)) = as" by simp with merge_list_length show "length (merge_list [] (map (\x. [x]) as)) \ length as" by metis qed lemma extract_IPSets_length: "length (extract_IPSets rs) \ 2 * length rs" apply(simp add: extract_IPSets_def) using extract_src_dst_ips_length mergesort_remdups_length by (metis add.right_neutral list.size(3)) (* export_code extract_IPSets in SML why you no work? *) lemma extract_equi0: "set (map wordinterval_to_set (extract_IPSets_generic0 sel rs)) = (\(base,len). ipset_from_cidr base len) ` sel ` match_sel ` set rs" proof(induction rs) case (Cons r rs) thus ?case apply(cases r, simp) using wordinterval_to_set_ipcidr_tuple_to_wordinterval by fastforce qed(simp) lemma src_ipPart_motivation: fixes rs defines "X \ (\(base,len). ipset_from_cidr base len) ` src ` match_sel ` set rs" assumes "\A \ X. B \ A \ B \ A = {}" and "s1 \ B" and "s2 \ B" shows "simple_fw rs (p\p_src:=s1\) = simple_fw rs (p\p_src:=s2\)" proof - have "\A \ (\(base,len). ipset_from_cidr base len) ` src ` match_sel ` set rs. B \ A \ B \ A = {} \ ?thesis" proof(induction rs) case Nil thus ?case by simp next case (Cons r rs) { fix m from \s1 \ B\ \s2 \ B\ have "B \ (case src m of (x, xa) \ ipset_from_cidr x xa) \ B \ (case src m of (x, xa) \ ipset_from_cidr x xa) = {} \ simple_matches m (p\p_src := s1\) \ simple_matches m (p\p_src := s2\)" apply(cases m) apply(rename_tac iiface oiface srca dst proto sports dports) apply(case_tac srca) apply(simp add: simple_matches.simps) by blast } note helper=this from Cons[simplified] show ?case apply(cases r, rename_tac m a) apply(simp) apply(case_tac a) using helper apply force+ done qed with assms show ?thesis by blast qed lemma src_ipPart: assumes "ipPartition (set (map wordinterval_to_set (extract_IPSets_generic0 src rs))) A" "B \ A" "s1 \ B" "s2 \ B" shows "simple_fw rs (p\p_src:=s1\) = simple_fw rs (p\p_src:=s2\)" proof - from src_ipPart_motivation[OF _ assms(3) assms(4)] have "\A \ (\(base,len). ipset_from_cidr base len) ` src ` match_sel ` set rs. B \ A \ B \ A = {} \ simple_fw rs (p\p_src:=s1\) = simple_fw rs (p\p_src:=s2\)" by fast thus ?thesis using assms(1) assms(2) unfolding ipPartition_def by (metis (full_types) Int_commute extract_equi0) qed (*basically a copy of src_ipPart*) lemma dst_ipPart: assumes "ipPartition (set (map wordinterval_to_set (extract_IPSets_generic0 dst rs))) A" "B \ A" "s1 \ B" "s2 \ B" shows "simple_fw rs (p\p_dst:=s1\) = simple_fw rs (p\p_dst:=s2\)" proof - have "\A \ (\(base,len). ipset_from_cidr base len) ` dst ` match_sel ` set rs. B \ A \ B \ A = {} \ simple_fw rs (p\p_dst:=s1\) = simple_fw rs (p\p_dst:=s2\)" proof(induction rs) case Nil thus ?case by simp next case (Cons r rs) { fix m from \s1 \ B\ \s2 \ B\ have "B \ (case dst m of (x, xa) \ ipset_from_cidr x xa) \ B \ (case dst m of (x, xa) \ ipset_from_cidr x xa) = {} \ simple_matches m (p\p_dst := s1\) \ simple_matches m (p\p_dst := s2\)" apply(cases m) apply(rename_tac iiface oiface src dsta proto sports dports) apply(case_tac dsta) apply(simp add: simple_matches.simps) by blast } note helper=this from Cons show ?case apply(simp) apply(case_tac r, rename_tac m a) apply(simp) apply(case_tac a) using helper apply force+ done qed thus ?thesis using assms(1) assms(2) unfolding ipPartition_def by (metis (full_types) Int_commute extract_equi0) qed (* OPTIMIZED PARTITIONING *) definition wordinterval_list_to_set :: "'a::len wordinterval list \ 'a::len word set" where "wordinterval_list_to_set ws = \(set (map wordinterval_to_set ws))" lemma wordinterval_list_to_set_compressed: "wordinterval_to_set (wordinterval_compress (foldr wordinterval_union xs Empty_WordInterval)) = wordinterval_list_to_set xs" proof(induction xs) qed(simp_all add: wordinterval_compress wordinterval_list_to_set_def) fun partIps :: "'a::len wordinterval \ 'a::len wordinterval list \ 'a::len wordinterval list" where "partIps _ [] = []" | "partIps s (t#ts) = (if wordinterval_empty s then (t#ts) else (if wordinterval_empty (wordinterval_intersection s t) then (t#(partIps s ts)) else (if wordinterval_empty (wordinterval_setminus t s) then (t#(partIps (wordinterval_setminus s t) ts)) else (wordinterval_intersection t s)#((wordinterval_setminus t s)# (partIps (wordinterval_setminus s t) ts)))))" lemma "partIps (WordInterval (1::ipv4addr) 1) [WordInterval 0 1] = [WordInterval 1 1, WordInterval 0 0]" by eval lemma partIps_length: "length (partIps s ts) \ (length ts) * 2" proof(induction ts arbitrary: s) case Cons thus ?case apply(simp) using le_Suc_eq by blast qed(simp) fun partitioningIps :: "'a::len wordinterval list \ 'a::len wordinterval list \ 'a::len wordinterval list" where "partitioningIps [] ts = ts" | "partitioningIps (s#ss) ts = partIps s (partitioningIps ss ts)" lemma partitioningIps_length: "length (partitioningIps ss ts) \ (2^length ss) * length ts" proof(induction ss) case Nil thus ?case by simp next case (Cons s ss) have "length (partIps s (partitioningIps ss ts)) \ length (partitioningIps ss ts) * 2" using partIps_length by fast with Cons show ?case by force qed lemma partIps_equi: "map wordinterval_to_set (partIps s ts) = partList4 (wordinterval_to_set s) (map wordinterval_to_set ts)" proof(induction ts arbitrary: s) qed(simp_all) lemma partitioningIps_equi: "map wordinterval_to_set (partitioningIps ss ts) = (partitioning1 (map wordinterval_to_set ss) (map wordinterval_to_set ts))" apply(induction ss arbitrary: ts) apply(simp; fail) apply(simp add: partIps_equi) done definition getParts :: "'i::len simple_rule list \ 'i wordinterval list" where "getParts rs = partitioningIps (extract_IPSets rs) [wordinterval_UNIV]" lemma partitioningIps_foldr: "partitioningIps ss ts = foldr partIps ss ts" by(induction ss) (simp_all) lemma getParts_foldr: "getParts rs = foldr partIps (extract_IPSets rs) [wordinterval_UNIV]" by(simp add: getParts_def partitioningIps_foldr) lemma getParts_length: "length (getParts rs) \ 2^(2 * length rs)" proof - from partitioningIps_length[where ss="extract_IPSets rs" and ts="[wordinterval_UNIV]"] have 1: "length (partitioningIps (extract_IPSets rs) [wordinterval_UNIV]) \ 2 ^ length (extract_IPSets rs)" by simp from extract_IPSets_length have "(2::nat) ^ length (extract_IPSets rs) \ 2 ^ (2 * length rs)" by simp with 1 have "length (partitioningIps (extract_IPSets rs) [wordinterval_UNIV]) \ 2 ^ (2 * length rs)" by linarith thus ?thesis by(simp add: getParts_def) qed lemma getParts_ipPartition: "ipPartition (set (map wordinterval_to_set (extract_IPSets rs))) (set (map wordinterval_to_set (getParts rs)))" proof - have hlp_rule: "{} \ set (map wordinterval_to_set ts) \ disjoint_list (map wordinterval_to_set ts) \ (wordinterval_list_to_set ss) \ (wordinterval_list_to_set ts) \ ipPartition (set (map wordinterval_to_set ss)) (set (map wordinterval_to_set (partitioningIps ss ts)))" for ts ss::"'a wordinterval list" by (metis ipPartitioning_helper_opt partitioningIps_equi wordinterval_list_to_set_def) have "disjoint_list [UNIV]" by(simp add: disjoint_list_def disjoint_def) have "ipPartition (set (map wordinterval_to_set ss)) (set (map wordinterval_to_set (partitioningIps ss [wordinterval_UNIV])))" for ss::"'a wordinterval list" apply(rule hlp_rule) apply(simp_all add: wordinterval_list_to_set_def \disjoint_list [UNIV]\) done thus ?thesis unfolding getParts_def by blast qed lemma getParts_complete: "wordinterval_list_to_set (getParts rs) = UNIV" proof - have "{} \ set (map wordinterval_to_set ts) \ (wordinterval_list_to_set ss) \ (wordinterval_list_to_set ts) \ wordinterval_list_to_set (partitioningIps ss ts) = (wordinterval_list_to_set ts)" for ss ts::"'a wordinterval list" using complete_helper by (metis partitioningIps_equi wordinterval_list_to_set_def) hence "wordinterval_list_to_set (getParts rs) = wordinterval_list_to_set [wordinterval_UNIV]" unfolding getParts_def by(simp add: wordinterval_list_to_set_def) also have "\ = UNIV" by (simp add: wordinterval_list_to_set_def) finally show ?thesis . qed theorem getParts_samefw: assumes "A \ set (map wordinterval_to_set (getParts rs))" "s1 \ A" "s2 \ A" shows "simple_fw rs (p\p_src:=s1\) = simple_fw rs (p\p_src:=s2\) \ simple_fw rs (p\p_dst:=s1\) = simple_fw rs (p\p_dst:=s2\)" proof - let ?X="(set (map wordinterval_to_set (getParts rs)))" from getParts_ipPartition have "ipPartition (set (map wordinterval_to_set (extract_IPSets rs))) ?X" . hence "ipPartition (set (map wordinterval_to_set (extract_IPSets_generic0 src rs))) ?X \ ipPartition (set (map wordinterval_to_set (extract_IPSets_generic0 dst rs))) ?X" by(simp add: extract_IPSets ipPartitionUnion image_Un) thus ?thesis using assms dst_ipPart src_ipPart by blast qed lemma partIps_nonempty: "ts \ [] \ partIps s ts \ []" by(induction ts arbitrary: s) simp_all lemma partitioningIps_nonempty: "ts \ [] \ partitioningIps ss ts \ []" proof(induction ss arbitrary: ts) case Nil thus ?case by simp next case (Cons s ss) thus ?case apply(cases ts) apply(simp; fail) apply(simp) using partIps_nonempty by blast qed (* lemma partIps_nonempty: "\t \ set ts. \ wordinterval_empty t \ {} \ set (map wordinterval_to_set (partIps s ts))" apply(induction ts arbitrary: s) apply(simp; fail) apply(simp) by blast *) lemma getParts_nonempty: "getParts rs \ []" by(simp add: getParts_def partitioningIps_nonempty) lemma getParts_nonempty_elems: "\w\set (getParts rs). \ wordinterval_empty w" unfolding getParts_def proof - have partitioning_nonempty: "\t \ set ts. \ wordinterval_empty t \ {} \ set (map wordinterval_to_set (partitioningIps ss ts))" for ts ss::"'a wordinterval list" proof(induction ss arbitrary: ts) case Nil thus ?case by auto case Cons thus ?case by (simp add: partIps_equi partList4_empty) qed have "\t \ set [wordinterval_UNIV].\ wordinterval_empty t" by(simp) with partitioning_nonempty have "{} \ set (map wordinterval_to_set (partitioningIps (extract_IPSets rs) [wordinterval_UNIV]))" by blast thus "\w\set (partitioningIps (extract_IPSets rs) [wordinterval_UNIV]). \ wordinterval_empty w" by auto qed (* HELPER FUNCTIONS UNIFICATION *) fun getOneIp :: "'a::len wordinterval \ 'a::len word" where "getOneIp (WordInterval b _) = b" | "getOneIp (RangeUnion r1 r2) = (if wordinterval_empty r1 then getOneIp r2 else getOneIp r1)" lemma getOneIp_elem: "\ wordinterval_empty W \ wordinterval_element (getOneIp W) W" by (induction W) simp_all record parts_connection = pc_iiface :: string pc_oiface :: string pc_proto :: primitive_protocol pc_sport :: "16 word" pc_dport :: "16 word" (* SAME FW DEFINITIONS AND PROOFS *) definition same_fw_behaviour :: "\<^cancel>\'pkt_ext itself \\ 'i::len word \ 'i word \ 'i simple_rule list \ bool" where "same_fw_behaviour \<^cancel>\TYPE('pkt_ext)\ a b rs \ \(p:: 'i::len simple_packet). simple_fw rs (p\p_src:=a\) = simple_fw rs (p\p_src:=b\) \ simple_fw rs (p\p_dst:=a\) = simple_fw rs (p\p_dst:=b\)" lemma getParts_same_fw_behaviour: "A \ set (map wordinterval_to_set (getParts rs)) \ s1 \ A \ s2 \ A \ same_fw_behaviour s1 s2 rs" unfolding same_fw_behaviour_def using getParts_samefw by blast definition "runFw s d c rs = simple_fw rs \p_iiface=pc_iiface c,p_oiface=pc_oiface c, p_src=s,p_dst=d, p_proto=pc_proto c, p_sport=pc_sport c,p_dport=pc_dport c, p_tcp_flags={TCP_SYN}, p_payload=''''\" text\We use @{const runFw} for executable code, but in general, everything applies to generic packets\ definition runFw_scheme :: "'i::len word \ 'i word \ 'b parts_connection_scheme \ ('i, 'a) simple_packet_scheme \ 'i simple_rule list \ state" where "runFw_scheme s d c p rs = simple_fw rs (p\p_iiface:=pc_iiface c, p_oiface:=pc_oiface c, p_src:=s, p_dst:=d, p_proto:=pc_proto c, p_sport:=pc_sport c, p_dport:=pc_dport c\)" lemma runFw_scheme: "runFw s d c rs = runFw_scheme s d c p rs" apply(simp add: runFw_def runFw_scheme_def) apply(case_tac p) apply(simp) apply(thin_tac _, simp) proof(induction rs) case Nil thus ?case by(simp; fail) next case(Cons r rs) obtain m a where r: "r = SimpleRule m a" by(cases r) simp from simple_matches_extended_packet[symmetric, of _ "pc_iiface c" "pc_oiface c" s d "pc_proto c" "pc_sport c" "pc_dport c" _ _ _ "{TCP_SYN}" "[]"] have pext: "simple_matches m \p_iiface = pc_iiface c, p_oiface = pc_oiface c, p_src = s, p_dst = d, p_proto = pc_proto c, p_sport = pc_sport c, p_dport = pc_dport c, p_tcp_flags = tcp_flags2, p_payload = payload2, \ = aux\ = simple_matches m \p_iiface = pc_iiface c, p_oiface = pc_oiface c, p_src = s, p_dst = d, p_proto = pc_proto c, p_sport = pc_sport c, p_dport = pc_dport c, p_tcp_flags = {TCP_SYN}, p_payload = []\" for tcp_flags2 payload2 and aux::'c by fast show ?case apply(simp add: r, cases a, simp) using Cons.IH by(simp add: pext)+ qed lemma has_default_policy_runFw: "has_default_policy rs \ runFw s d c rs = Decision FinalAllow \ runFw s d c rs = Decision FinalDeny" by(simp add: runFw_def has_default_policy) definition same_fw_behaviour_one :: "'i::len word \ 'i word \ 'a parts_connection_scheme \ 'i simple_rule list \ bool" where "same_fw_behaviour_one ip1 ip2 c rs \ \d s. runFw ip1 d c rs = runFw ip2 d c rs \ runFw s ip1 c rs = runFw s ip2 c rs" lemma same_fw_spec: "same_fw_behaviour ip1 ip2 rs \ same_fw_behaviour_one ip1 ip2 c rs" apply(simp add: same_fw_behaviour_def same_fw_behaviour_one_def runFw_def) apply(rule conjI) apply(clarify) apply(erule_tac x="\p_iiface = pc_iiface c, p_oiface = pc_oiface c, p_src = ip1, p_dst= d, p_proto = pc_proto c, p_sport = pc_sport c, p_dport = pc_dport c, p_tcp_flags = {TCP_SYN}, p_payload=''''\" in allE) apply(simp;fail) apply(clarify) apply(erule_tac x="\p_iiface = pc_iiface c, p_oiface = pc_oiface c, p_src = s, p_dst= ip1, p_proto = pc_proto c, p_sport = pc_sport c, p_dport = pc_dport c, p_tcp_flags = {TCP_SYN}, p_payload=''''\" in allE) apply(simp) done text\Is an equivalence relation\ lemma same_fw_behaviour_one_equi: "same_fw_behaviour_one x x c rs" "same_fw_behaviour_one x y c rs = same_fw_behaviour_one y x c rs" "same_fw_behaviour_one x y c rs \ same_fw_behaviour_one y z c rs \ same_fw_behaviour_one x z c rs" unfolding same_fw_behaviour_one_def by metis+ lemma same_fw_behaviour_equi: "same_fw_behaviour x x rs" "same_fw_behaviour x y rs = same_fw_behaviour y x rs" "same_fw_behaviour x y rs \ same_fw_behaviour y z rs \ same_fw_behaviour x z rs" unfolding same_fw_behaviour_def by auto lemma runFw_sameFw_behave: fixes W :: "'i::len word set set" shows "\A \ W. \a1 \ A. \a2 \ A. same_fw_behaviour_one a1 a2 c rs \ \ W = UNIV \ \B \ W. \b \ B. runFw ip1 b c rs = runFw ip2 b c rs \ \B \ W. \b \ B. runFw b ip1 c rs = runFw b ip2 c rs \ same_fw_behaviour_one ip1 ip2 c rs" proof - assume a1: "\A \ W. \a1 \ A. \a2 \ A. same_fw_behaviour_one a1 a2 c rs" and a2: "\ W = UNIV " and a3: "\B \ W. \b \ B. runFw ip1 b c rs = runFw ip2 b c rs" and a4: "\B \ W. \b \ B. runFw b ip1 c rs = runFw b ip2 c rs" have relation_lem: "\D \ W. \d1 \ D. \d2 \ D. \s. f s d1 = f s d2 \ \ W = UNIV \ \B \ W. \b \ B. f s1 b = f s2 b \ f s1 d = f s2 d" for W and f::"'c \ 'b \ 'd" and s1 d s2 by (metis UNIV_I Union_iff) from a1 have a1':"\A\W. \a1\A. \a2\A. \s. runFw s a1 c rs = runFw s a2 c rs" unfolding same_fw_behaviour_one_def by fast from relation_lem[OF a1' a2 a3] have s1: "\ d. runFw ip1 d c rs = runFw ip2 d c rs" by simp from a1 have a1'':"\A\W. \a1\A. \a2\A. \d. runFw a1 d c rs = runFw a2 d c rs" unfolding same_fw_behaviour_one_def by fast from relation_lem[OF a1'' a2 a4] have s2: "\ s. runFw s ip1 c rs = runFw s ip2 c rs" by simp from s1 s2 show "same_fw_behaviour_one ip1 ip2 c rs" unfolding same_fw_behaviour_one_def by fast qed lemma sameFw_behave_sets: "\w\set A. \a1 \ w. \a2 \ w. same_fw_behaviour_one a1 a2 c rs \ \w1\set A. \w2\set A. \a1\w1. \a2\w2. same_fw_behaviour_one a1 a2 c rs \ \w1\set A. \w2\set A. \a1\w1. \a2\w2. same_fw_behaviour_one a1 a2 c rs" proof - assume a1: "\w\set A. \a1 \ w. \a2 \ w. same_fw_behaviour_one a1 a2 c rs" and "\w1\set A. \w2\set A. \a1\w1. \a2\w2. same_fw_behaviour_one a1 a2 c rs" from this have "\w1\set A. \w2\set A. \a1\w1. \a2\w2. same_fw_behaviour_one a1 a2 c rs" using same_fw_behaviour_one_equi(3) by metis from a1 this show "\w1\set A. \w2\set A. \a1\w1. \a2\w2. same_fw_behaviour_one a1 a2 c rs" using same_fw_behaviour_one_equi(3) by metis qed definition groupWIs :: "parts_connection \ 'i::len simple_rule list \ 'i wordinterval list list" where "groupWIs c rs = (let W = getParts rs in (let f = (\wi. (map (\d. runFw (getOneIp wi) d c rs) (map getOneIp W), map (\s. runFw s (getOneIp wi) c rs) (map getOneIp W))) in groupF f W))" lemma groupWIs_not_empty: "groupWIs c rs \ []" proof - have "getParts rs \ []" by(simp add: getParts_def partitioningIps_nonempty) with groupF_empty have "\f. groupF f (getParts rs) \ []" by blast thus ?thesis by(simp add: groupWIs_def Let_def) blast qed lemma groupWIs_not_empty_elem: "V \ set (groupWIs c rs) \ V \ []" by(simp add: groupWIs_def Let_def groupF_empty_elem) lemma groupWIs_not_empty_elems: assumes V: "V \ set (groupWIs c rs)" and w: "w \ set V" shows "\ wordinterval_empty w" proof - have "\w\set (concat (groupWIs c rs)). \ wordinterval_empty w" apply(subst groupWIs_def) apply(subst Let_def)+ apply(subst groupF_concat_set) using getParts_nonempty_elems by blast from this V w show ?thesis by auto qed lemma groupParts_same_fw_wi0: assumes "V \ set (groupWIs c rs)" shows "\w \ set (map wordinterval_to_set V). \a1 \ w. \a2 \ w. same_fw_behaviour_one a1 a2 c rs" proof - have "\A\set (map wordinterval_to_set (concat (groupWIs c rs))). \a1\A. \a2\A. same_fw_behaviour_one a1 a2 c rs" apply(subst groupWIs_def) apply(subst Let_def)+ apply(subst set_map) apply(subst groupF_concat_set) using getParts_same_fw_behaviour same_fw_spec by fastforce from this assms show ?thesis by force qed lemma groupWIs_same_fw_not: "A \ set (groupWIs c rs) \ B \ set (groupWIs c rs) \ A \ B \ \aw \ set (map wordinterval_to_set A). \bw \ set (map wordinterval_to_set B). \a \ aw. \b \ bw. \ same_fw_behaviour_one a b c rs" proof - assume asm: "A \ set (groupWIs c rs)" "B \ set (groupWIs c rs)" "A \ B" from this have b1: "\aw \ set A. \bw \ set B. (map (\d. runFw (getOneIp aw) d c rs) (map getOneIp (getParts rs)), map (\s. runFw s (getOneIp aw) c rs) (map getOneIp (getParts rs))) \ (map (\d. runFw (getOneIp bw) d c rs) (map getOneIp (getParts rs)), map (\s. runFw s (getOneIp bw) c rs) (map getOneIp (getParts rs)))" apply(simp add: groupWIs_def Let_def) using groupF_nequality by fastforce have same_behave_runFw_not: "(map (\d. runFw x1 d c rs) W, map (\s. runFw s x1 c rs) W) \ (map (\d. runFw x2 d c rs) W, map (\s. runFw s x2 c rs) W) \ \ same_fw_behaviour_one x1 x2 c rs" for x1 x2 W by (simp add: same_fw_behaviour_one_def) (blast) have "\C \ set (groupWIs c rs). \c \ set C. getOneIp c \ wordinterval_to_set c" apply(simp add: groupWIs_def Let_def) using getParts_nonempty_elems groupF_set getOneIp_elem by fastforce from this b1 asm have "\aw \ set (map wordinterval_to_set A). \bw \ set (map wordinterval_to_set B). \a \ aw. \b \ bw. (map (\d. runFw a d c rs) (map getOneIp (getParts rs)), map (\s. runFw s a c rs) (map getOneIp (getParts rs))) \ (map (\d. runFw b d c rs) (map getOneIp (getParts rs)), map (\s. runFw s b c rs) (map getOneIp (getParts rs)))" by (simp) (blast) from this same_behave_runFw_not asm have " \aw \ set (map wordinterval_to_set A). \bw \ set (map wordinterval_to_set B). \a \ aw. \b \ bw. \ same_fw_behaviour_one a b c rs" by fast from this groupParts_same_fw_wi0[of A c rs] groupParts_same_fw_wi0[of B c rs] asm have "\aw \ set (map wordinterval_to_set A). \bw \ set (map wordinterval_to_set B). \a \ aw. \b \ bw. \ same_fw_behaviour_one a b c rs" apply(simp) using same_fw_behaviour_one_equi(3) by blast from this groupParts_same_fw_wi0[of A c rs] groupParts_same_fw_wi0[of B c rs] asm show "\aw \ set (map wordinterval_to_set A). \bw \ set (map wordinterval_to_set B). \a \ aw. \b \ bw. \ same_fw_behaviour_one a b c rs" apply(simp) using same_fw_behaviour_one_equi(3) by fast qed (*beginning is copy&paste of previous proof*) lemma groupParts_same_fw_wi1: "V \ set (groupWIs c rs) \ \w1 \ set V. \w2 \ set V. \a1 \ wordinterval_to_set w1. \a2 \ wordinterval_to_set w2. same_fw_behaviour_one a1 a2 c rs" proof - assume asm: "V \ set (groupWIs c rs)" from getParts_same_fw_behaviour same_fw_spec have b1: "\A \ set (map wordinterval_to_set (getParts rs)) . \a1 \ A. \a2 \ A. same_fw_behaviour_one a1 a2 c rs" by fast from getParts_complete have complete: "\(set (map wordinterval_to_set (getParts rs))) = UNIV" by(simp add: wordinterval_list_to_set_def) from getParts_nonempty_elems have nonempty: "\w\set (getParts rs). \ wordinterval_empty w" by simp { fix W x1 x2 assume a1: "\A \ set (map wordinterval_to_set W). \a1 \ A. \a2 \ A. same_fw_behaviour_one a1 a2 c rs" and a2: "wordinterval_list_to_set W = UNIV" and a3: "\w \ set W. \ wordinterval_empty w" and a4: "(map (\d. runFw x1 d c rs) (map getOneIp W), map (\s. runFw s x1 c rs) (map getOneIp W)) = (map (\d. runFw x2 d c rs) (map getOneIp W), map (\s. runFw s x2 c rs) (map getOneIp W))" from a3 a4 getOneIp_elem have b1: "\B \ set (map wordinterval_to_set W). \b \ B. runFw x1 b c rs = runFw x2 b c rs" by fastforce from a3 a4 getOneIp_elem have b2: "\B \ set (map wordinterval_to_set W). \b \ B. runFw b x1 c rs = runFw b x2 c rs" by fastforce from runFw_sameFw_behave[OF a1 _ b1 b2] a2[unfolded wordinterval_list_to_set_def] have "same_fw_behaviour_one x1 x2 c rs" by simp } note same_behave_runFw=this from same_behave_runFw[OF b1 getParts_complete nonempty] groupF_equality[of "(\wi. (map (\d. runFw (getOneIp wi) d c rs) (map getOneIp (getParts rs)), map (\s. runFw s (getOneIp wi) c rs) (map getOneIp (getParts rs))))" "(getParts rs)"] asm have b2: "\a1\set V. \a2\set V. same_fw_behaviour_one (getOneIp a1) (getOneIp a2) c rs" apply(subst (asm) groupWIs_def) apply(subst (asm) Let_def)+ by fast from groupWIs_not_empty_elems asm have "\w \ set V. \ wordinterval_empty w" by blast from this b2 getOneIp_elem have b3: "\w1\set (map wordinterval_to_set V). \w2\set (map wordinterval_to_set V). \ip1\ w1. \ip2\w2. same_fw_behaviour_one ip1 ip2 c rs" by (simp) (blast) from groupParts_same_fw_wi0 asm have "\A\set (map wordinterval_to_set V). \a1\ A. \a2\ A. same_fw_behaviour_one a1 a2 c rs" by metis from sameFw_behave_sets[OF this b3] show "\w1 \ set V. \w2 \ set V. \a1 \ wordinterval_to_set w1. \a2 \ wordinterval_to_set w2. same_fw_behaviour_one a1 a2 c rs" by force qed lemma groupParts_same_fw_wi2: "V \ set (groupWIs c rs) \ \ip1 \ wordinterval_list_to_set V. \ip2 \ wordinterval_list_to_set V. same_fw_behaviour_one ip1 ip2 c rs" using groupParts_same_fw_wi0 groupParts_same_fw_wi1 apply (simp add: wordinterval_list_to_set_def) by fast lemma groupWIs_same_fw_not2: "A \ set (groupWIs c rs) \ B \ set (groupWIs c rs) \ A \ B \ \ip1 \ wordinterval_list_to_set A. \ip2 \ wordinterval_list_to_set B. \ same_fw_behaviour_one ip1 ip2 c rs" apply(simp add: wordinterval_list_to_set_def) using groupWIs_same_fw_not by fastforce (*I like this version -- corny*) lemma "A \ set (groupWIs c rs) \ B \ set (groupWIs c rs) \ \ip1 \ wordinterval_list_to_set A. \ip2 \ wordinterval_list_to_set B. same_fw_behaviour_one ip1 ip2 c rs \ A = B" using groupWIs_same_fw_not2 by blast lemma groupWIs_complete: "(\x\ set (groupWIs c rs). wordinterval_list_to_set x) = (UNIV::'i::len word set)" proof - have "(\ y \ (\x\ set (groupWIs c rs). set x). wordinterval_to_set y) = (UNIV::'i word set)" apply(simp add: groupWIs_def Let_def groupF_Union_set) using getParts_complete wordinterval_list_to_set_def by fastforce thus ?thesis by(simp add: wordinterval_list_to_set_def) qed (*begin groupWIs1 and groupWIs2 optimization*) definition groupWIs1 :: "'a parts_connection_scheme \ 'i::len simple_rule list \ 'i wordinterval list list" where "groupWIs1 c rs = (let P = getParts rs in (let W = map getOneIp P in (let f = (\wi. (map (\d. runFw (getOneIp wi) d c rs) W, map (\s. runFw s (getOneIp wi) c rs) W)) in map (map fst) (groupF snd (map (\x. (x, f x)) P)))))" lemma groupWIs_groupWIs1_equi: "groupWIs1 c rs = groupWIs c rs" apply(subst groupWIs1_def) apply(subst groupWIs_def) using groupF_tuple by metis definition simple_conn_matches :: "'i::len simple_match \ parts_connection \ bool" where "simple_conn_matches m c \ (match_iface (iiface m) (pc_iiface c)) \ (match_iface (oiface m) (pc_oiface c)) \ (match_proto (proto m) (pc_proto c)) \ (simple_match_port (sports m) (pc_sport c)) \ (simple_match_port (dports m) (pc_dport c))" lemma simple_conn_matches_simple_match_any: "simple_conn_matches simple_match_any c" apply (simp add: simple_conn_matches_def simple_match_any_def match_ifaceAny) - apply (subgoal_tac "(65535::16 word) = max_word") + apply (subgoal_tac "(65535::16 word) = - 1") apply (simp only:) apply simp_all done lemma has_default_policy_simple_conn_matches: "has_default_policy rs \ has_default_policy [r\rs . simple_conn_matches (match_sel r) c]" apply(induction rs rule: has_default_policy.induct) apply(simp; fail) apply(simp add: simple_conn_matches_simple_match_any; fail) apply(simp) apply(intro conjI) apply(simp split: if_split_asm; fail) apply(simp add: has_default_policy_fst split: if_split_asm) done lemma filter_conn_fw_lem: "runFw s d c (filter (\r. simple_conn_matches (match_sel r) c) rs) = runFw s d c rs" apply(simp add: runFw_def simple_conn_matches_def match_sel_def) apply(induction rs "\p_iiface = pc_iiface c, p_oiface = pc_oiface c, p_src = s, p_dst = d, p_proto = pc_proto c, p_sport = pc_sport c, p_dport = pc_dport c, p_tcp_flags = {TCP_SYN},p_payload=''''\" rule: simple_fw.induct) apply(simp add: simple_matches.simps)+ done (*performance: despite optimization, this function takes quite long and can be optimized*) definition groupWIs2 :: "parts_connection \ 'i::len simple_rule list \ 'i wordinterval list list" where "groupWIs2 c rs = (let P = getParts rs in (let W = map getOneIp P in (let filterW = (filter (\r. simple_conn_matches (match_sel r) c) rs) in (let f = (\wi. (map (\d. runFw (getOneIp wi) d c filterW) W, map (\s. runFw s (getOneIp wi) c filterW) W)) in map (map fst) (groupF snd (map (\x. (x, f x)) P))))))" lemma groupWIs1_groupWIs2_equi: "groupWIs2 c rs = groupWIs1 c rs" by(simp add: groupWIs2_def groupWIs1_def filter_conn_fw_lem) lemma groupWIs_code[code]: "groupWIs c rs = groupWIs2 c rs" using groupWIs1_groupWIs2_equi groupWIs_groupWIs1_equi by metis (*end groupWIs1 and groupWIs2 optimization*) (*begin groupWIs3 optimization*) fun matching_dsts :: "'i::len word \ 'i simple_rule list \ 'i wordinterval \ 'i wordinterval" where "matching_dsts _ [] _ = Empty_WordInterval" | "matching_dsts s ((SimpleRule m Accept)#rs) acc_dropped = (if simple_match_ip (src m) s then wordinterval_union (wordinterval_setminus (ipcidr_tuple_to_wordinterval (dst m)) acc_dropped) (matching_dsts s rs acc_dropped) else matching_dsts s rs acc_dropped)" | "matching_dsts s ((SimpleRule m Drop)#rs) acc_dropped = (if simple_match_ip (src m) s then matching_dsts s rs (wordinterval_union (ipcidr_tuple_to_wordinterval (dst m)) acc_dropped) else matching_dsts s rs acc_dropped)" lemma matching_dsts_pull_out_accu: "wordinterval_to_set (matching_dsts s rs (wordinterval_union a1 a2)) = wordinterval_to_set (matching_dsts s rs a2) - wordinterval_to_set a1" apply(induction s rs a2 arbitrary: a1 a2 rule: matching_dsts.induct) apply(simp_all) by blast+ (*a copy of matching_dsts*) fun matching_srcs :: "'i::len word \ 'i simple_rule list \ 'i wordinterval \ 'i wordinterval" where "matching_srcs _ [] _ = Empty_WordInterval" | "matching_srcs d ((SimpleRule m Accept)#rs) acc_dropped = (if simple_match_ip (dst m) d then wordinterval_union (wordinterval_setminus (ipcidr_tuple_to_wordinterval (src m)) acc_dropped) (matching_srcs d rs acc_dropped) else matching_srcs d rs acc_dropped)" | "matching_srcs d ((SimpleRule m Drop)#rs) acc_dropped = (if simple_match_ip (dst m) d then matching_srcs d rs (wordinterval_union (ipcidr_tuple_to_wordinterval (src m)) acc_dropped) else matching_srcs d rs acc_dropped)" lemma matching_srcs_pull_out_accu: "wordinterval_to_set (matching_srcs d rs (wordinterval_union a1 a2)) = wordinterval_to_set (matching_srcs d rs a2) - wordinterval_to_set a1" apply(induction d rs a2 arbitrary: a1 a2 rule: matching_srcs.induct) apply(simp_all) by blast+ lemma matching_dsts: "\r \ set rs. simple_conn_matches (match_sel r) c \ wordinterval_to_set (matching_dsts s rs Empty_WordInterval) = {d. runFw s d c rs = Decision FinalAllow}" proof(induction rs) case Nil thus ?case by (simp add: runFw_def) next case (Cons r rs) obtain m a where r: "r = SimpleRule m a" by(cases r, blast) from Cons.prems r have simple_match_ip_Accept: "\d. simple_match_ip (src m) s \ runFw s d c (SimpleRule m Accept # rs) = Decision FinalAllow \ simple_match_ip (dst m) d \ runFw s d c rs = Decision FinalAllow" by(simp add: simple_conn_matches_def runFw_def simple_matches.simps) { fix d a have "\ simple_match_ip (src m) s \ runFw s d c (SimpleRule m a # rs) = Decision FinalAllow \ runFw s d c rs = Decision FinalAllow" apply(cases a) by(simp add: simple_conn_matches_def runFw_def simple_matches.simps)+ } note not_simple_match_ip=this from Cons.prems r have simple_match_ip_Drop: "\d. simple_match_ip (src m) s \ runFw s d c (SimpleRule m Drop # rs) = Decision FinalAllow \ \ simple_match_ip (dst m) d \ runFw s d c rs = Decision FinalAllow" by(simp add: simple_conn_matches_def runFw_def simple_matches.simps) show ?case proof(cases a) case Accept with r Cons show ?thesis apply(simp, intro conjI impI) apply(simp add: simple_match_ip_Accept wordinterval_to_set_ipcidr_tuple_to_wordinterval_simple_match_ip_set) apply blast apply(simp add: not_simple_match_ip; fail) done next case Drop with r Cons show ?thesis apply(simp,intro conjI impI) apply(simp add: simple_match_ip_Drop matching_dsts_pull_out_accu wordinterval_to_set_ipcidr_tuple_to_wordinterval_simple_match_ip_set) apply blast apply(simp add: not_simple_match_ip; fail) done qed qed lemma matching_srcs: "\r \ set rs. simple_conn_matches (match_sel r) c \ wordinterval_to_set (matching_srcs d rs Empty_WordInterval) = {s. runFw s d c rs = Decision FinalAllow}" proof(induction rs) case Nil thus ?case by (simp add: runFw_def) next case (Cons r rs) obtain m a where r: "r = SimpleRule m a" by(cases r, blast) from Cons.prems r have simple_match_ip_Accept: "\s. simple_match_ip (dst m) d \ runFw s d c (SimpleRule m Accept # rs) = Decision FinalAllow \ simple_match_ip (src m) s \ runFw s d c rs = Decision FinalAllow" by(simp add: simple_conn_matches_def runFw_def simple_matches.simps) { fix s a have "\ simple_match_ip (dst m) d \ runFw s d c (SimpleRule m a # rs) = Decision FinalAllow \ runFw s d c rs = Decision FinalAllow" apply(cases a) by(simp add: simple_conn_matches_def runFw_def simple_matches.simps)+ } note not_simple_match_ip=this from Cons.prems r have simple_match_ip_Drop: "\s. simple_match_ip (dst m) d \ runFw s d c (SimpleRule m Drop # rs) = Decision FinalAllow \ \ simple_match_ip (src m) s \ runFw s d c rs = Decision FinalAllow" by(simp add: simple_conn_matches_def runFw_def simple_matches.simps) show ?case proof(cases a) case Accept with r Cons show ?thesis apply(simp, intro conjI impI) apply(simp add: simple_match_ip_Accept wordinterval_to_set_ipcidr_tuple_to_wordinterval_simple_match_ip_set) apply blast apply(simp add: not_simple_match_ip; fail) done next case Drop with r Cons show ?thesis apply(simp,intro conjI impI) apply(simp add: simple_match_ip_Drop matching_srcs_pull_out_accu wordinterval_to_set_ipcidr_tuple_to_wordinterval_simple_match_ip_set) apply blast apply(simp add: not_simple_match_ip; fail) done qed qed (*TODO: if we can get wordinterval_element to log runtime ...*) definition groupWIs3_default_policy :: "parts_connection \ 'i::len simple_rule list \ 'i wordinterval list list" where "groupWIs3_default_policy c rs = (let P = getParts rs in (let W = map getOneIp P in (let filterW = (filter (\r. simple_conn_matches (match_sel r) c) rs) in (let f = (\wi. let mtch_dsts = (matching_dsts (getOneIp wi) filterW Empty_WordInterval); mtch_srcs = (matching_srcs (getOneIp wi) filterW Empty_WordInterval) in (map (\d. wordinterval_element d mtch_dsts) W, map (\s. wordinterval_element s mtch_srcs) W)) in map (map fst) (groupF snd (map (\x. (x, f x)) P))))))" lemma groupWIs3_default_policy_groupWIs2: fixes rs :: "'i::len simple_rule list" assumes "has_default_policy rs" shows "groupWIs2 c rs = groupWIs3_default_policy c rs" proof - { fix filterW s d from matching_dsts[where c=c] have "filterW = filter (\r. simple_conn_matches (match_sel r) c) rs \ wordinterval_element d (matching_dsts s filterW Empty_WordInterval) \ runFw s d c filterW = Decision FinalAllow" by force } note matching_dsts_filterW=this[simplified] { fix filterW s d from matching_srcs[where c=c] have "filterW = filter (\r. simple_conn_matches (match_sel r) c) rs \ wordinterval_element s (matching_srcs d filterW Empty_WordInterval) \ runFw s d c filterW = Decision FinalAllow" by force } note matching_srcs_filterW=this[simplified] { fix W and rs :: "'i::len simple_rule list" assume assms': "has_default_policy rs" have "groupF (\wi. (map (\d. runFw (getOneIp wi) d c rs = Decision FinalAllow) (map getOneIp W), map (\s. runFw s (getOneIp wi) c rs = Decision FinalAllow) (map getOneIp W))) W = groupF (\wi. (map (\d. runFw (getOneIp wi) d c rs) (map getOneIp W), map (\s. runFw s (getOneIp wi) c rs) (map getOneIp W))) W" proof - { (*unused fresh generic variables. 'a is used for the tuple already*) fix f1::"'w \ 'u \ 'v" and f2::" 'w \ 'u \ 'x" and x and y and g1::"'w \ 'u \ 'y" and g2::"'w \ 'u \ 'z" and W::"'u list" assume 1: "\w \ set W. (f1 x) w = (f1 y) w \ (f2 x) w = (f2 y) w" and 2: "\w \ set W. (g1 x) w = (g1 y) w \ (g2 x) w = (g2 y) w" have " ((map (f1 x) W, map (g1 x) W) = (map (f1 y) W, map (g1 y) W)) \ ((map (f2 x) W, map (g2 x) W) = (map (f2 y) W, map (g2 y) W))" proof - from 1 have p1: "(map (f1 x) W = map (f1 y) W \ map (f2 x) W = map (f2 y) W)" by(induction W)(simp_all) from 2 have p2: "(map (g1 x) W = map (g1 y) W \ map (g2 x) W = map (g2 y) W)" by(induction W)(simp_all) from p1 p2 show ?thesis by fast qed } note map_over_tuples_equal_helper=this show ?thesis apply(rule groupF_cong) apply(intro ballI) apply(rule map_over_tuples_equal_helper) using has_default_policy_runFw[OF assms'] by metis+ qed } note has_default_policy_groupF=this[simplified] from assms show ?thesis apply(simp add: groupWIs3_default_policy_def groupWIs_code[symmetric]) apply(subst groupF_tuple[symmetric]) apply(simp add: Let_def) apply(simp add: matching_srcs_filterW matching_dsts_filterW) apply(subst has_default_policy_groupF) apply(simp add: has_default_policy_simple_conn_matches; fail) apply(simp add: groupWIs_def Let_def filter_conn_fw_lem) done qed definition groupWIs3 :: "parts_connection \ 'i::len simple_rule list \ 'i wordinterval list list" where "groupWIs3 c rs = (if has_default_policy rs then groupWIs3_default_policy c rs else groupWIs2 c rs)" lemma groupWIs3: "groupWIs3 = groupWIs" by(simp add: fun_eq_iff groupWIs3_def groupWIs_code groupWIs3_default_policy_groupWIs2) (*end groupWIs3 optimization*) (*construct partitions. main function!*) definition build_ip_partition :: "parts_connection \ 'i::len simple_rule list \ 'i wordinterval list" where "build_ip_partition c rs = map (\xs. wordinterval_sort (wordinterval_compress (foldr wordinterval_union xs Empty_WordInterval))) (groupWIs3 c rs)" theorem build_ip_partition_same_fw: "V \ set (build_ip_partition c rs) \ \ip1::'i::len word \ wordinterval_to_set V. \ip2::'i::len word \ wordinterval_to_set V. same_fw_behaviour_one ip1 ip2 c rs" apply(simp add: build_ip_partition_def groupWIs3) using wordinterval_list_to_set_compressed groupParts_same_fw_wi2 wordinterval_sort by blast theorem build_ip_partition_same_fw_min: "A \ set (build_ip_partition c rs) \ B \ set (build_ip_partition c rs) \ A \ B \ \ip1::'i::len word \ wordinterval_to_set A. \ip2::'i::len word \ wordinterval_to_set B. \ same_fw_behaviour_one ip1 ip2 c rs" apply(simp add: build_ip_partition_def groupWIs3) using groupWIs_same_fw_not2 wordinterval_list_to_set_compressed wordinterval_sort by blast theorem build_ip_partition_complete: "(\x\set (build_ip_partition c rs). wordinterval_to_set x) = (UNIV :: 'i::len word set)" proof - have "wordinterval_to_set (foldr wordinterval_union x Empty_WordInterval) = \(set (map wordinterval_to_set x))" for x::"'i wordinterval list" by(induction x) simp_all thus ?thesis apply(simp add: build_ip_partition_def groupWIs3 wordinterval_compress wordinterval_sort) using groupWIs_complete[simplified wordinterval_list_to_set_def] by simp qed lemma build_ip_partition_no_empty_elems: "wi \ set (build_ip_partition c rs) \ \ wordinterval_empty wi" proof - assume "wi \ set (build_ip_partition c rs)" hence assm: "wi \ (\xs. wordinterval_sort (wordinterval_compress (foldr wordinterval_union xs Empty_WordInterval))) ` set (groupWIs c rs)" by(simp add: build_ip_partition_def groupWIs3) from assm obtain wi_orig where 1: "wi_orig \ set (groupWIs c rs)" and 2: "wi = wordinterval_sort (wordinterval_compress (foldr wordinterval_union wi_orig Empty_WordInterval))" by blast from 1 groupWIs_not_empty_elem have i1: "wi_orig \ []" by blast from 1 groupWIs_not_empty_elems have i2: "\w. w \ set wi_orig \ \ wordinterval_empty w" by simp from i1 i2 have "wordinterval_to_set (foldr wordinterval_union wi_orig Empty_WordInterval) \ {}" by(induction wi_orig) simp_all with 2 show ?thesis by(simp add: wordinterval_compress wordinterval_sort) qed lemma build_ip_partition_disjoint: "V1 \ set (build_ip_partition c rs) \ V2 \ set (build_ip_partition c rs) \ V1 \ V2 \ wordinterval_to_set V1 \ wordinterval_to_set V2 = {}" by (metis build_ip_partition_same_fw build_ip_partition_same_fw_min disjoint_iff) lemma map_wordinterval_to_set_distinct: assumes distinct: "distinct xs" and disjoint: "(\x1 \ set xs. \x2 \ set xs. x1 \ x2 \ wordinterval_to_set x1 \ wordinterval_to_set x2 = {})" and notempty: "\x \ set xs. \ wordinterval_empty x" shows "distinct (map wordinterval_to_set xs)" proof - have "\ wordinterval_empty x1 \ wordinterval_to_set x1 \ wordinterval_to_set x2 = {} \ wordinterval_to_set x1 \ wordinterval_to_set x2" for x1::"('b::len) wordinterval" and x2 by auto with disjoint notempty have "(\x1 \ set xs. \x2 \ set xs. x1 \ x2 \ wordinterval_to_set x1 \ wordinterval_to_set x2)" by force with distinct show "distinct (map wordinterval_to_set xs)" proof(induction xs) case Cons thus ?case by simp fast qed(simp) qed lemma map_getOneIp_distinct: assumes distinct: "distinct xs" and disjoint: "(\x1 \ set xs. \x2 \ set xs. x1 \ x2 \ wordinterval_to_set x1 \ wordinterval_to_set x2 = {})" and notempty: "\x \ set xs. \ wordinterval_empty x" shows "distinct (map getOneIp xs)" proof - have "\ wordinterval_empty x \ \ wordinterval_empty xa \ wordinterval_to_set x \ wordinterval_to_set xa = {} \ getOneIp x \ getOneIp xa" for x xa::"'b::len wordinterval" by(fastforce dest: getOneIp_elem) with disjoint notempty have "(\x1 \ set xs. \x2 \ set xs. x1 \ x2 \ getOneIp x1 \ getOneIp x2)" by metis with distinct show ?thesis proof(induction xs) case Cons thus ?case by simp fast qed(simp) qed lemma getParts_disjoint_list: "disjoint_list (map wordinterval_to_set (getParts rs))" proof- have disjoint_list_partitioningIps: "{} \ set (map wordinterval_to_set ts) \ disjoint_list (map wordinterval_to_set ts) \ (wordinterval_list_to_set ss) \ (wordinterval_list_to_set ts) \ disjoint_list (map wordinterval_to_set (partitioningIps ss ts))" for ts::"'a::len wordinterval list" and ss by (simp add: partitioning1_disjoint_list partitioningIps_equi wordinterval_list_to_set_def) have "{} \ set (map wordinterval_to_set [wordinterval_UNIV])" and "disjoint_list (map wordinterval_to_set [wordinterval_UNIV])" and "wordinterval_list_to_set (extract_IPSets rs) \ wordinterval_list_to_set [wordinterval_UNIV]" by(simp add: wordinterval_list_to_set_def disjoint_list_def disjoint_def)+ thus ?thesis unfolding getParts_def by(rule disjoint_list_partitioningIps) qed lemma build_ip_partition_distinct: "distinct (map wordinterval_to_set (build_ip_partition c rs))" proof - have "(wordinterval_to_set \ (\xs. wordinterval_sort (wordinterval_compress (foldr wordinterval_union xs Empty_WordInterval)))) ws = \(set (map wordinterval_to_set ws))" for ws::"'a::len wordinterval list" proof(induction ws) qed(simp_all add: wordinterval_compress wordinterval_sort) hence hlp1: "map wordinterval_to_set (build_ip_partition c rs) = map (\x. \(set (map wordinterval_to_set x))) (groupWIs c rs)" unfolding build_ip_partition_def groupWIs3 by auto \ \generic rule\ have "\x \ set xs. \ wordinterval_empty x \ disjoint_list (map wordinterval_to_set xs) \ distinct (map (\x. \(set (map wordinterval_to_set x))) (groupF f xs))" for f::"'x::len wordinterval \ 'y" and xs::"'x::len wordinterval list" proof(induction f xs rule: groupF.induct) case 1 thus ?case by simp next case (2 f x xs) have hlp_internal: "\ (set (map (\x. \(set (map wordinterval_to_set x))) (groupF f xs))) = \ (set (map wordinterval_to_set xs))" for f::"'x wordinterval \ 'y" and xs by(induction f xs rule: groupF.induct) (auto) from 2(2,3) have "wordinterval_to_set x \ \(wordinterval_to_set ` set xs) = {}" by(auto simp add: disjoint_def disjoint_list_def) hence "\ (wordinterval_to_set x) \ \(wordinterval_to_set ` set xs)" using 2(2) by auto hence "\ wordinterval_to_set x \ \(set (map wordinterval_to_set [y\xs . f x \ f y]))" by auto hence "\ wordinterval_to_set x \ (\x\{xa \ set xs. f x = f xa}. wordinterval_to_set x) \ \ (set (map (\x. \(set (map wordinterval_to_set x))) (groupF f [y\xs . f x \ f y])))" unfolding hlp_internal by blast hence g1: "wordinterval_to_set x \ (\x\{xa \ set xs. f x = f xa}. wordinterval_to_set x) \ (\x. \x\set x. wordinterval_to_set x) ` set (groupF f [y\xs . f x \ f y])" by force from 2(3) have "distinct (map wordinterval_to_set [y\xs . f x \ f y])" by (simp add: disjoint_list_def distinct_map_filter) moreover from 2 have "disjoint (wordinterval_to_set ` {xa \ set xs. f x \ f xa})" by(simp add: disjoint_def disjoint_list_def) ultimately have g2: "distinct (map (\x. \x\set x. wordinterval_to_set x) (groupF f [y\xs . f x \ f y]))" using 2(1,2) unfolding disjoint_list_def by(simp) from g1 g2 show ?case by simp qed with getParts_disjoint_list getParts_nonempty_elems have "distinct (map (\x. \(set (map wordinterval_to_set x))) (groupF (\wi. (map (\d. runFw (getOneIp wi) d c rs) (map getOneIp (getParts rs)), map (\s. runFw s (getOneIp wi) c rs) (map getOneIp (getParts rs)))) (getParts rs)))" by blast thus ?thesis unfolding hlp1 groupWIs_def Let_def by presburger qed lemma build_ip_partition_distinct': "distinct (build_ip_partition c rs)" using build_ip_partition_distinct distinct_mapI by blast subsection\Service Matrix over an IP Address Space Partition\ definition simple_firewall_without_interfaces :: "'i::len simple_rule list \ bool" where "simple_firewall_without_interfaces rs \ \m \ match_sel ` set rs. iiface m = ifaceAny \ oiface m = ifaceAny" lemma simple_fw_no_interfaces: assumes no_ifaces: "simple_firewall_without_interfaces rs" shows "simple_fw rs p = simple_fw rs (p\ p_iiface:= x, p_oiface:= y\)" proof - from no_ifaces have "\r\set rs. iiface (match_sel r) = ifaceAny \ oiface (match_sel r) = ifaceAny" by(simp add: simple_firewall_without_interfaces_def) thus ?thesis apply(induction rs p rule:simple_fw.induct) by(simp_all add: simple_matches.simps match_ifaceAny) qed lemma runFw_no_interfaces: assumes no_ifaces: "simple_firewall_without_interfaces rs" shows "runFw s d c rs = runFw s d (c\ pc_iiface:= x, pc_oiface:= y\) rs" apply(simp add: runFw_def) apply(subst simple_fw_no_interfaces[OF no_ifaces]) by(simp) lemma[code_unfold]: "simple_firewall_without_interfaces rs \ \m \ set rs. iiface (match_sel m) = ifaceAny \ oiface (match_sel m) = ifaceAny" by(simp add: simple_firewall_without_interfaces_def) (*only defined for simple_firewall_without_interfaces*) definition access_matrix :: "parts_connection \ 'i::len simple_rule list \ ('i word \ 'i wordinterval) list \ ('i word \ 'i word) list" where "access_matrix c rs \ (let W = build_ip_partition c rs; R = map getOneIp W in (zip R W, [(s, d)\all_pairs R. runFw s d c rs = Decision FinalAllow]))" lemma access_matrix_nodes_defined: "(V,E) = access_matrix c rs \ (s, d) \ set E \ s \ dom (map_of V)" and "(V,E) = access_matrix c rs \ (s, d) \ set E \ d \ dom (map_of V)" by(auto simp add: access_matrix_def Let_def all_pairs_def) text\For all the entries @{term E} of the matrix, the access is allowed\ lemma "(V,E) = access_matrix c rs \ (s, d) \ set E \ runFw s d c rs = Decision FinalAllow" by(auto simp add: access_matrix_def Let_def) text\However, the entries are only a representation of a whole set of IP addresses. For all IP addresses which the entries represent, the access must be allowed.\ (*TODO: move to generic library*) lemma map_of_zip_map: "map_of (zip (map f rs) rs) k = Some v \ k = f v" apply(induction rs) apply(simp) apply(simp split: if_split_asm) done lemma access_matrix_sound: assumes matrix: "(V,E) = access_matrix c rs" and repr: "(s_repr, d_repr) \ set E" and s_range: "(map_of V) s_repr = Some s_range" and s: "s \ wordinterval_to_set s_range" and d_range: "(map_of V) d_repr = Some d_range" and d: "d \ wordinterval_to_set d_range" shows "runFw s d c rs = Decision FinalAllow" proof - let ?part="(build_ip_partition c rs)" have V: "V = (zip (map getOneIp ?part) ?part)" using matrix by(simp add: access_matrix_def Let_def) (*have "E = [(s, d)\all_pairs (map getOneIp ?part). runFw s d c rs = Decision FinalAllow]" using matrix by(simp add: access_matrix_def Let_def) with repr have "(s_repr, d_repr) \ set (all_pairs (map getOneIp ?part))" by simp hence "s_repr \ set (map getOneIp ?part)" and "d_repr \ set (map getOneIp ?part)" by(simp add: all_pairs_set)+*) (*from s_range have "(s_repr, s_range) \ set V" by (simp add: map_of_SomeD)*) from matrix repr have repr_Allow: "runFw s_repr d_repr c rs = Decision FinalAllow" by(auto simp add: access_matrix_def Let_def) have s_range_in_part: "s_range \ set ?part" using V s_range by (fastforce elim: in_set_zipE dest: map_of_SomeD) with build_ip_partition_no_empty_elems have "\ wordinterval_empty s_range" by simp have d_range_in_part: "d_range \ set ?part" using V d_range by (fastforce elim: in_set_zipE dest: map_of_SomeD) with build_ip_partition_no_empty_elems have "\ wordinterval_empty d_range" by simp from map_of_zip_map V s_range have "s_repr = getOneIp s_range" by fast with \\ wordinterval_empty s_range\ getOneIp_elem wordinterval_element_set_eq have "s_repr \ wordinterval_to_set s_range" by blast from map_of_zip_map V d_range have "d_repr = getOneIp d_range" by fast with \\ wordinterval_empty d_range\ getOneIp_elem wordinterval_element_set_eq have "d_repr \ wordinterval_to_set d_range" by blast from s_range_in_part have s_range_in_part': "s_range \ set (build_ip_partition c rs)" by simp from d_range_in_part have d_range_in_part': "d_range \ set (build_ip_partition c rs)" by simp from build_ip_partition_same_fw[OF s_range_in_part', unfolded same_fw_behaviour_one_def] s \s_repr \ wordinterval_to_set s_range\ have "\d. runFw s_repr d c rs = runFw s d c rs" by blast with repr_Allow have 1: "runFw s d_repr c rs = Decision FinalAllow" by simp from build_ip_partition_same_fw[OF d_range_in_part', unfolded same_fw_behaviour_one_def] d \d_repr \ wordinterval_to_set d_range\ have "\s. runFw s d_repr c rs = runFw s d c rs" by blast with 1 have 2: "runFw s d c rs = Decision FinalAllow" by simp thus ?thesis . qed lemma distinct_map_getOneIp_obtain: "v \ set xs \ distinct (map getOneIp xs) \ \s_repr. map_of (zip (map getOneIp xs) xs) s_repr = Some v" proof(induction xs) case Nil thus ?case by simp next case (Cons x xs) consider "v = x" | "v \ set xs" using Cons.prems(1) by fastforce thus ?case proof(cases) case 1 thus ?thesis by simp blast next case 2 with Cons.IH Cons.prems(2) obtain s_repr where s_repr: "map_of (zip (map getOneIp xs) xs) s_repr = Some v" by force show ?thesis proof(cases "s_repr \ getOneIp x") case True with Cons.prems s_repr show ?thesis by(rule_tac x=s_repr in exI, simp) next case False with Cons.prems s_repr show ?thesis by(fastforce elim: in_set_zipE) qed qed qed lemma access_matrix_complete: fixes rs :: "'i::len simple_rule list" assumes matrix: "(V,E) = access_matrix c rs" and allow: "runFw s d c rs = Decision FinalAllow" shows "\s_repr d_repr s_range d_range. (s_repr, d_repr) \ set E \ (map_of V) s_repr = Some s_range \ s \ wordinterval_to_set s_range \ (map_of V) d_repr = Some d_range \ d \ wordinterval_to_set d_range" proof - let ?part="(build_ip_partition c rs)" have V: "V = zip (map getOneIp ?part) ?part" using matrix by(simp add: access_matrix_def Let_def) have E: "E = [(s, d)\all_pairs (map getOneIp ?part). runFw s d c rs = Decision FinalAllow]" using matrix by(simp add: access_matrix_def Let_def) have build_ip_partition_obtain: "\V. V \ set (build_ip_partition c rs) \ s \ wordinterval_to_set V" for s using build_ip_partition_complete by blast have distinct_map_getOneIp_build_ip_partition_obtain: "v \ set (build_ip_partition c rs) \ \s_repr. map_of (zip (map getOneIp (build_ip_partition c rs)) (build_ip_partition c rs)) s_repr = Some v" for v and rs :: "'i::len simple_rule list" proof(erule distinct_map_getOneIp_obtain) show "distinct (map getOneIp (build_ip_partition c rs))" apply(rule map_getOneIp_distinct) subgoal using build_ip_partition_distinct' by blast subgoal using build_ip_partition_disjoint build_ip_partition_distinct' by blast subgoal using build_ip_partition_no_empty_elems[simplified] by auto done qed from build_ip_partition_obtain obtain s_range where "s_range \ set ?part" and "s \ wordinterval_to_set s_range" by blast from this distinct_map_getOneIp_build_ip_partition_obtain V obtain s_repr where ex_s1: "(map_of V) s_repr = Some s_range" and ex_s2: "s \ wordinterval_to_set s_range" by blast from build_ip_partition_obtain obtain d_range where "d_range \ set ?part" and "d \ wordinterval_to_set d_range" by blast from this distinct_map_getOneIp_build_ip_partition_obtain V obtain d_repr where ex_d1: "(map_of V) d_repr = Some d_range" and ex_d2: "d \ wordinterval_to_set d_range" by blast have 1: "s_repr \ getOneIp ` set (build_ip_partition c rs)" using V \map_of V s_repr = Some s_range\ by (fastforce elim: in_set_zipE dest: map_of_SomeD) have 2: "d_repr \ getOneIp ` set (build_ip_partition c rs)" using V \map_of V d_repr = Some d_range\ by (fastforce elim: in_set_zipE dest: map_of_SomeD) have "runFw s_repr d_repr c rs = Decision FinalAllow" proof - have f1: "(\w wa p ss. \ same_fw_behaviour_one w wa (p::parts_connection) ss \ (\wb wc. runFw w wb p ss = runFw wa wb p ss \ runFw wc w p ss = runFw wc wa p ss)) \ (\w wa p ss. (\wb wc. runFw w wb (p::parts_connection) ss \ runFw wa wb p ss \ runFw wc w p ss \ runFw wc wa p ss) \ same_fw_behaviour_one w wa p ss)" unfolding same_fw_behaviour_one_def by blast from \s_range \ set (build_ip_partition c rs)\ have f2: "same_fw_behaviour_one s s_repr c rs" by (metis (no_types) map_of_zip_map V build_ip_partition_no_empty_elems build_ip_partition_same_fw ex_s1 ex_s2 getOneIp_elem wordinterval_element_set_eq) from \d_range \ set (build_ip_partition c rs)\ have "same_fw_behaviour_one d_repr d c rs" by (metis (no_types) map_of_zip_map V build_ip_partition_no_empty_elems build_ip_partition_same_fw ex_d1 ex_d2 getOneIp_elem wordinterval_element_set_eq) with f1 f2 show ?thesis using allow by metis qed hence ex1: "(s_repr, d_repr) \ set E" by(simp add: E all_pairs_set 1 2) thus ?thesis using ex1 ex_s1 ex_s2 ex_d1 ex_d2 by blast qed theorem access_matrix: fixes rs :: "'i::len simple_rule list" assumes matrix: "(V,E) = access_matrix c rs" shows "(\s_repr d_repr s_range d_range. (s_repr, d_repr) \ set E \ (map_of V) s_repr = Some s_range \ s \ wordinterval_to_set s_range \ (map_of V) d_repr = Some d_range \ d \ wordinterval_to_set d_range) \ runFw s d c rs = Decision FinalAllow" using matrix access_matrix_sound access_matrix_complete by blast text\ For a @{typ "'i::len simple_rule list"} and a fixed @{typ parts_connection}, we support to partition the IP address space; for IP addresses of arbitrary length (eg., IPv4, IPv6). All members of a partition have the same access rights: @{thm build_ip_partition_same_fw [no_vars]} Minimal: @{thm build_ip_partition_same_fw_min [no_vars]} The resulting access control matrix is sound and complete: @{thm access_matrix [no_vars]} Theorem reads: For a fixed connection, you can look up IP addresses (source and destination pairs) in the matrix if and only if the firewall accepts this src,dst IP address pair for the fixed connection. Note: The matrix is actually a graph (nice visualization!), you need to look up IP addresses in the Vertices and check the access of the representants in the edges. If you want to visualize the graph (e.g. with Graphviz or tkiz): The vertices are the node description (i.e. header; @{term "dom V"} is the label for each node which will also be referenced in the edges, @{term "ran V"} is the human-readable description for each node (i.e. the full IP range it represents)), the edges are the edges. Result looks nice. Theorem also tells us that this visualization is correct. \ (*construct an ip partition and print it in some usable format returns: (vertices, edges) where vertices = (name, list of ip addresses this vertex corresponds to) and edges = (name \ name) list Note that the WordInterval is already sorted, which is important for prettyness! *) text\Only defined for @{const simple_firewall_without_interfaces}\ definition access_matrix_pretty_ipv4 :: "parts_connection \ 32 simple_rule list \ (string \ string) list \ (string \ string) list" where "access_matrix_pretty_ipv4 c rs \ if \ simple_firewall_without_interfaces rs then undefined else (let (V,E) = (access_matrix c rs); formatted_nodes = map (\(v_repr, v_range). (ipv4addr_toString v_repr, ipv4addr_wordinterval_toString v_range)) V; formatted_edges = map (\(s,d). (ipv4addr_toString s, ipv4addr_toString d)) E in (formatted_nodes, formatted_edges) )" definition access_matrix_pretty_ipv4_code :: "parts_connection \ 32 simple_rule list \ (string \ string) list \ (string \ string) list" where "access_matrix_pretty_ipv4_code c rs \ if \ simple_firewall_without_interfaces rs then undefined else (let W = build_ip_partition c rs; R = map getOneIp W; U = all_pairs R in (zip (map ipv4addr_toString R) (map ipv4addr_wordinterval_toString W), map (\(x,y). (ipv4addr_toString x, ipv4addr_toString y)) [(s, d)\all_pairs R. runFw s d c rs = Decision FinalAllow]))" lemma access_matrix_pretty_ipv4_code[code]: "access_matrix_pretty_ipv4 = access_matrix_pretty_ipv4_code" by(simp add: fun_eq_iff access_matrix_pretty_ipv4_def access_matrix_pretty_ipv4_code_def Let_def access_matrix_def map_prod_fun_zip) definition access_matrix_pretty_ipv6 :: "parts_connection \ 128 simple_rule list \ (string \ string) list \ (string \ string) list" where "access_matrix_pretty_ipv6 c rs \ if \ simple_firewall_without_interfaces rs then undefined else (let (V,E) = (access_matrix c rs); formatted_nodes = map (\(v_repr, v_range). (ipv6addr_toString v_repr, ipv6addr_wordinterval_toString v_range)) V; formatted_edges = map (\(s,d). (ipv6addr_toString s, ipv6addr_toString d)) E in (formatted_nodes, formatted_edges) )" definition access_matrix_pretty_ipv6_code :: "parts_connection \ 128 simple_rule list \ (string \ string) list \ (string \ string) list" where "access_matrix_pretty_ipv6_code c rs \ if \ simple_firewall_without_interfaces rs then undefined else (let W = build_ip_partition c rs; R = map getOneIp W; U = all_pairs R in (zip (map ipv6addr_toString R) (map ipv6addr_wordinterval_toString W), map (\(x,y). (ipv6addr_toString x, ipv6addr_toString y)) [(s, d)\all_pairs R. runFw s d c rs = Decision FinalAllow]))" lemma access_matrix_pretty_ipv6_code[code]: "access_matrix_pretty_ipv6 = access_matrix_pretty_ipv6_code" by(simp add: fun_eq_iff access_matrix_pretty_ipv6_def access_matrix_pretty_ipv6_code_def Let_def access_matrix_def map_prod_fun_zip) definition parts_connection_ssh where "parts_connection_ssh \ \pc_iiface=''1'', pc_oiface=''1'', pc_proto=TCP, pc_sport=10000, pc_dport=22\" definition parts_connection_http where "parts_connection_http \ \pc_iiface=''1'', pc_oiface=''1'', pc_proto=TCP, pc_sport=10000, pc_dport=80\" definition mk_parts_connection_TCP :: "16 word \ 16 word \ parts_connection" where "mk_parts_connection_TCP sport dport = \pc_iiface=''1'', pc_oiface=''1'', pc_proto=TCP, pc_sport=sport, pc_dport=dport\" lemma "mk_parts_connection_TCP 10000 22 = parts_connection_ssh" "mk_parts_connection_TCP 10000 80 = parts_connection_http" by(simp_all add: mk_parts_connection_TCP_def parts_connection_ssh_def parts_connection_http_def) value[code] "partitioningIps [WordInterval (0::ipv4addr) 0] [WordInterval 0 2, WordInterval 0 2]" text_raw\ Here is an example of a really large and complicated firewall: \begin{figure} \centering \resizebox{\linewidth}{!}{% \tikzset{every loop/.style={looseness=1}} \tikzset{myptr/.style={decoration={markings,mark=at position 1 with % {\arrow[scale=2,>=latex]{>}}},shorten >=0.1cm,shorten <=0.2cm, postaction={decorate}}}% \tikzset{myloop/.style={-to}}% \begin{tikzpicture} \node (m) at (2,-2) {$\{224.0.0.0 .. 239.255.255.255\}$}; \node[align=left] (inet) at (-3,8.5) {$\{0.0.0.0 .. 126.255.255.255\} \cup \{128.0.0.0 .. 131.159.13.255\} \cup $ \\ $ \{131.159.16.0 .. 131.159.19.255\} \cup \{131.159.22.0 .. 138.246.253.4\} \cup $ \\ $ \{138.246.253.11 .. 185.86.231.255\} \cup \{185.86.236.0 .. 188.1.239.85\} \cup $ \\ $ \{188.1.239.87 .. 188.95.232.63\} \cup \{188.95.232.224 .. 188.95.232.255\} \cup $\\ $ \{188.95.240.0 .. 192.48.106.255\} \cup \{192.48.108.0 .. 223.255.255.255\} \cup $\\ $ \{240.0.0.0 .. 255.255.255.255\}$}; \node[align=left] (internal) at (-5,-10) {$\{131.159.14.0 .. 131.159.14.7\} \cup \{131.159.14.12 .. 131.159.14.25\} \cup $ \\ $ 131.159.14.27 \cup \{131.159.14.29 .. 131.159.14.33\} \cup $ \\ $ \{131.159.14.38 .. 131.159.14.39\} \cup 131.159.14.41 \cup $ \\ $ \{131.159.14.43 .. 131.159.14.51\} \cup \{131.159.14.53 .. 131.159.14.55\} \cup $ \\ $ \{131.159.14.57 .. 131.159.14.59\} \cup \{131.159.14.61 .. 131.159.14.68\} \cup $ \\ $ 131.159.14.70 .. 131.159.14.82\} \cup \{131.159.14.84 .. 131.159.14.103\} \cup $ \\ $ \{131.159.14.105 .. 131.159.14.124\} \cup \{131.159.14.126 .. 131.159.14.136\} \cup $ \\ $ \{131.159.14.138 .. 131.159.14.139\} \cup \{131.159.14.141 .. 131.159.14.144\} \cup $ \\ $ \{131.159.14.147 .. 131.159.14.154\} \cup \{131.159.14.157 .. 131.159.14.162\} \cup $ \\ $ \{131.159.14.164 .. 131.159.14.168\} \cup \{131.159.14.170 .. 131.159.14.200\} \cup $ \\ $ \{131.159.14.202 .. 131.159.14.213\} \cup \{131.159.14.215 .. 131.159.15.4\} \cup $ \\ $ 131.159.15.6 \cup \{131.159.15.14 .. 131.159.15.15\} \cup $ \\ $ \{131.159.15.21 .. 131.159.15.22\} \cup 131.159.15.26 \cup 131.159.15.28 \cup $ \\ $ 131.159.15.30 \cup \{131.159.15.33 .. 131.159.15.35\} \cup $ \\ $ \{131.159.15.37 .. 131.159.15.38\} \cup 131.159.15.40 \cup $ \\ $ \{131.159.15.45 .. 131.159.15.46\} \cup \{131.159.15.48 .. 131.159.15.49\} \cup $ \\ $ \{131.159.15.52 .. 131.159.15.55\} \cup 131.159.15.57 \cup 131.159.15.59 \cup $ \\ $ \{131.159.15.61 .. 131.159.15.67\} \cup \{131.159.15.70 .. 131.159.15.196\} \cup $ \\ $ \{131.159.15.198 .. 131.159.15.227\} \cup \{131.159.15.229 .. 131.159.15.233\} \cup $ \\ $ \{131.159.15.235 .. 131.159.15.246\} \cup \{131.159.15.250 .. 131.159.15.255\} \cup $ \\ $ \{131.159.20.0 .. 131.159.20.20\} \cup \{131.159.20.22 .. 131.159.20.28\} \cup $ \\ $ \{131.159.20.30 .. 131.159.20.35\} \cup \{131.159.20.37 .. 131.159.20.44\} \cup $ \\ $ \{131.159.20.46 .. 131.159.20.51\} \cup \{131.159.20.53 .. 131.159.20.58\} \cup $ \\ $ \{131.159.20.60 .. 131.159.20.62\} \cup \{131.159.20.64 .. 131.159.20.70\} \cup $ \\ $ \{131.159.20.72 .. 131.159.20.73\} \cup \{131.159.20.75 .. 131.159.20.84\} \cup $ \\ $ \{131.159.20.86 .. 131.159.20.96\} \cup \{131.159.20.98 .. 131.159.20.119\} \cup $ \\ $ \{131.159.20.121 .. 131.159.20.138\} \cup \{131.159.20.140 .. 131.159.20.149\} \cup $ \\ $ \{131.159.20.152 .. 131.159.20.154\} \cup \{131.159.20.156 .. 131.159.20.159\} \cup $ \\ $ \{131.159.20.161 .. 131.159.20.164\} \cup \{131.159.20.167 .. 131.159.20.179\} \cup $ \\ $ \{131.159.20.181 .. 131.159.20.184\} \cup \{131.159.20.186 .. 131.159.20.199\} \cup $ \\ $ \{131.159.20.201 .. 131.159.20.232\} \cup \{131.159.20.235 .. 131.159.20.255\} \cup $ \\ $ \{185.86.232.0 .. 185.86.235.255\} \cup \{188.95.233.0 .. 188.95.233.3\} \cup $ \\ $ \{188.95.233.5 .. 188.95.233.8\} \cup \{188.95.233.10 .. 188.95.233.255\} \cup $ \\ $ \{192.48.107.0 .. 192.48.107.255\}$}; \node[align=left] (srvs) at (10,0) {$\{131.159.14.8 .. 131.159.14.11\} \cup 131.159.14.26 \cup 131.159.14.28 \cup $ \\ $ \{131.159.14.34 .. 131.159.14.37\} \cup 131.159.14.40 \cup 131.159.14.42 \cup $ \\ $ 131.159.14.52 \cup 131.159.14.56 \cup 131.159.14.60 \cup 131.159.14.69 \cup $ \\ $ 131.159.14.83 \cup 131.159.14.104 \cup 131.159.14.125 \cup 131.159.14.137 \cup $ \\ $ 131.159.14.140 \cup \{131.159.14.145 .. 131.159.14.146\} \cup $ \\ $ \{131.159.14.155 .. 131.159.14.156\} \cup 131.159.14.163 \cup 131.159.14.169 \cup $ \\ $ 131.159.14.201 \cup 131.159.14.214 \cup 131.159.15.5 \cup $ \\ $ \{131.159.15.7 .. 131.159.15.13\} \cup \{131.159.15.16 .. 131.159.15.20\} \cup $ \\ $ \{131.159.15.23 .. 131.159.15.25\} \cup 131.159.15.27 \cup 131.159.15.29 \cup $ \\ $ \{131.159.15.31 .. 131.159.15.32\} \cup 131.159.15.36 \cup 131.159.15.39 \cup $ \\ $ \{131.159.15.41 .. 131.159.15.44\} \cup 131.159.15.47 \cup 131.159.15.51 \cup $ \\ $ 131.159.15.56 \cup 131.159.15.58 \cup 131.159.15.60 \cup $ \\ $ \{131.159.15.68 .. 131.159.15.69\} \cup 131.159.15.197 \cup 131.159.15.228 \cup $ \\ $ 131.159.15.234 \cup \{131.159.15.247 .. 131.159.15.249\} \cup 131.159.20.21 \cup $ \\ $ 131.159.20.29 \cup 131.159.20.36 \cup 131.159.20.45 \cup 131.159.20.52 \cup $ \\ $ 131.159.20.59 \cup 131.159.20.63 \cup 131.159.20.71 \cup 131.159.20.74 \cup $ \\ $ 131.159.20.85 \cup 131.159.20.97 \cup 131.159.20.120 \cup 131.159.20.139 \cup $ \\ $ \{131.159.20.150 .. 131.159.20.151\} \cup 131.159.20.155 \cup 131.159.20.160 \cup $ \\ $ \{131.159.20.165 .. 131.159.20.166\} \cup 131.159.20.180 \cup 131.159.20.185 \cup $ \\ $ 131.159.20.200 \cup \{131.159.20.233 .. 131.159.20.234\} \cup $ \\ $ \{131.159.21.0 .. 131.159.21.255\} \cup \{188.95.232.192 .. 188.95.232.223\} \cup $ \\ $ 188.95.233.4 \cup 188.95.233.9 \cup \{188.95.234.0 .. 188.95.239.255\}$}; \node[align=left] (ips1) at (-3,-1) {$188.1.239.86 \cup \{188.95.232.64 .. 188.95.232.191\}$}; \node[align=left] (ips2) at (10,-8) {$\{138.246.253.6 .. 138.246.253.10\}$}; \node[align=left] (ip3) at (5,-6) {$138.246.253.5$}; \node[align=left] (ip4) at (4.5,-8) {$131.159.15.50$}; \node[align=left] (l) at (8,-10) {$\{127.0.0.0 .. 127.255.255.255\}$}; \draw[myloop] (m) to[loop above] (m); \draw[myptr] (m) to (srvs); \draw[myptr] (inet) to (m); \draw[myptr] (inet) to (srvs); \draw[myptr,shorten <=-0.8cm] (internal) to (m); \draw[myptr] (internal) to (inet); \draw[myloop] (internal) to[loop above] (internal); \draw[myptr] (internal) to (srvs); \draw[myptr] (internal) to (ips1); \draw[myptr] (internal) to (ips2); \draw[myptr] (internal) to (ip3); \draw[myptr] (internal) to (ip4); \draw[myptr] (internal) to (l); \draw[myptr] (srvs) to (m); \draw[myptr] (srvs) to (inet); \draw[myptr] (srvs) to (internal); \draw[myloop] (srvs) to[loop above] (srvs); \draw[myptr] (srvs) to (ips1); \draw[myptr] (srvs) to (ips2); \draw[myptr] (srvs) to (ip3); \draw[myptr] (srvs) to (ip4); \draw[myptr] (srvs) to (l); \draw[myptr] (ips1) to (m); \draw[myptr] (ips1) to (inet); \draw[myptr] (ips1) to (internal); \draw[myptr] (ips1) to (srvs); \draw[myloop] (ips1.west) to[loop left] (ips1); \draw[myptr] (ips1) to (ips2); \draw[myptr] (ips1) to (ip3); \draw[myptr] (ips1) to (ip4); \draw[myptr] (ips1) to (l); \draw[myptr] (ips2) to (m); \draw[myptr] (ips2) to (srvs); \draw[myptr] (ips2) to (ip4); \draw[myptr] (ip3) to (m); \draw[myptr] (ip3) to (internal); \draw[myptr] (ip3) to (srvs); \draw[myptr] (ip3) to (ip4); \draw[myptr] (ip4) to (m); \draw[myptr] (ip4) to (inet); \draw[myptr] (ip4) to (internal); \draw[myptr] (ip4) to (srvs); \draw[myptr] (ip4) to (ips1); \draw[myptr] (ip4) to (ips2); \draw[myptr] (ip4) to (ip3); \draw[myloop] (ip4) to[loop below] (ip4); \draw[myptr] (ip4) to (l); \end{tikzpicture}% } \caption{TUM ssh Service Matrix} \label{fig:tumssh} \end{figure} \ end diff --git a/thys/Simple_Firewall/SimpleFw_Semantics.thy b/thys/Simple_Firewall/SimpleFw_Semantics.thy --- a/thys/Simple_Firewall/SimpleFw_Semantics.thy +++ b/thys/Simple_Firewall/SimpleFw_Semantics.thy @@ -1,423 +1,423 @@ section\Simple Firewall Semantics\ theory SimpleFw_Semantics imports SimpleFw_Syntax IP_Addresses.IP_Address IP_Addresses.Prefix_Match begin fun simple_match_ip :: "('i::len word \ nat) \ 'i::len word \ bool" where "simple_match_ip (base, len) p_ip \ p_ip \ ipset_from_cidr base len" lemma wordinterval_to_set_ipcidr_tuple_to_wordinterval_simple_match_ip_set: "wordinterval_to_set (ipcidr_tuple_to_wordinterval ip) = {d. simple_match_ip ip d}" proof - { fix s and d :: "'a::len word \ nat" from wordinterval_to_set_ipcidr_tuple_to_wordinterval have "s \ wordinterval_to_set (ipcidr_tuple_to_wordinterval d) \ simple_match_ip d s" by(cases d) auto } thus ?thesis by blast qed \ \by the way, the words do not wrap around\ lemma "{(253::8 word) .. 8} = {}" by simp fun simple_match_port :: "(16 word \ 16 word) \ 16 word \ bool" where "simple_match_port (s,e) p_p \ p_p \ {s..e}" fun simple_matches :: "'i::len simple_match \ ('i, 'a) simple_packet_scheme \ bool" where "simple_matches m p \ (match_iface (iiface m) (p_iiface p)) \ (match_iface (oiface m) (p_oiface p)) \ (simple_match_ip (src m) (p_src p)) \ (simple_match_ip (dst m) (p_dst p)) \ (match_proto (proto m) (p_proto p)) \ (simple_match_port (sports m) (p_sport p)) \ (simple_match_port (dports m) (p_dport p))" text\The semantics of a simple firewall: just iterate over the rules sequentially\ fun simple_fw :: "'i::len simple_rule list \ ('i, 'a) simple_packet_scheme \ state" where "simple_fw [] _ = Undecided" | "simple_fw ((SimpleRule m Accept)#rs) p = (if simple_matches m p then Decision FinalAllow else simple_fw rs p)" | "simple_fw ((SimpleRule m Drop)#rs) p = (if simple_matches m p then Decision FinalDeny else simple_fw rs p)" fun simple_fw_alt where "simple_fw_alt [] _ = Undecided" | "simple_fw_alt (r#rs) p = (if simple_matches (match_sel r) p then (case action_sel r of Accept \ Decision FinalAllow | Drop \ Decision FinalDeny) else simple_fw_alt rs p)" lemma simple_fw_alt: "simple_fw r p = simple_fw_alt r p" by(induction rule: simple_fw.induct) simp_all definition simple_match_any :: "'i::len simple_match" where "simple_match_any \ \iiface=ifaceAny, oiface=ifaceAny, src=(0,0), dst=(0,0), proto=ProtoAny, sports=(0,65535), dports=(0,65535) \" lemma simple_match_any: "simple_matches simple_match_any p" proof - - have *: "(65535::16 word) = max_word" + have *: "(65535::16 word) = - 1" by simp show ?thesis by (simp add: simple_match_any_def ipset_from_cidr_0 match_ifaceAny *) qed text\we specify only one empty port range\ definition simple_match_none :: "'i::len simple_match" where "simple_match_none \ \iiface=ifaceAny, oiface=ifaceAny, src=(1,0), dst=(0,0), proto=ProtoAny, sports=(1,0), dports=(0,65535) \" lemma simple_match_none: "\ simple_matches simple_match_none p" proof - show ?thesis by(simp add: simple_match_none_def) qed fun empty_match :: "'i::len simple_match \ bool" where "empty_match \iiface=_, oiface=_, src=_, dst=_, proto=_, sports=(sps1, sps2), dports=(dps1, dps2) \ \ (sps1 > sps2) \ (dps1 > dps2)" lemma empty_match: "empty_match m \ (\(p::('i::len, 'a) simple_packet_scheme). \ simple_matches m p)" proof assume "empty_match m" thus "\p. \ simple_matches m p" by(cases m) fastforce next assume assm: "\(p::('i::len, 'a) simple_packet_scheme). \ simple_matches m p" obtain iif oif sip dip protocol sps1 sps2 dps1 dps2 where m: "m = \iiface = iif, oiface = oif, src = sip, dst = dip, proto = protocol, sports = (sps1, sps2), dports = (dps1, dps2)\" by(cases m) force show "empty_match m" proof(simp add: m) let ?x="\p. dps1 \ p_dport p \ p_sport p \ sps2 \ sps1 \ p_sport p \ match_proto protocol (p_proto p) \ simple_match_ip dip (p_dst p) \ simple_match_ip sip (p_src p) \ match_iface oif (p_oiface p) \ match_iface iif (p_iiface p) \ \ p_dport p \ dps2" from assm have nomatch: "\(p::('i::len, 'a) simple_packet_scheme). ?x p" by(simp add: m) { fix ips::"'i::len word \ nat" have "a \ ipset_from_cidr a n" for a::"'i::len word" and n using ipset_from_cidr_lowest by auto hence "simple_match_ip ips (fst ips)" by(cases ips) simp } note ips=this have proto: "match_proto protocol (case protocol of ProtoAny \ TCP | Proto p \ p)" by(simp split: protocol.split) { fix ifce have "match_iface ifce (iface_sel ifce)" by (simp add: match_iface_eqI) } note ifaces=this { fix p::"('i, 'a) simple_packet_scheme" from nomatch have "?x p" by blast } note pkt1=this obtain p::"('i, 'a) simple_packet_scheme" where [simp]: "p_iiface p = iface_sel iif" "p_oiface p = iface_sel oif" "p_src p = fst sip" "p_dst p = fst dip" "p_proto p = (case protocol of ProtoAny \ TCP | Proto p \ p)" "p_sport p = sps1" "p_dport p = dps1" by (meson simple_packet.select_convs) note pkt=pkt1[of p, simplified] from pkt ips proto ifaces have " sps1 \ sps2 \ \ dps1 \ dps2" by blast thus "sps2 < sps1 \ dps2 < dps1" by fastforce qed qed lemma nomatch: "\ simple_matches m p \ simple_fw (SimpleRule m a # rs) p = simple_fw rs p" by(cases a, simp_all del: simple_matches.simps) subsection\Simple Ports\ fun simpl_ports_conjunct :: "(16 word \ 16 word) \ (16 word \ 16 word) \ (16 word \ 16 word)" where "simpl_ports_conjunct (p1s, p1e) (p2s, p2e) = (max p1s p2s, min p1e p2e)" lemma "{(p1s:: 16 word) .. p1e} \ {p2s .. p2e} = {max p1s p2s .. min p1e p2e}" by(simp) lemma simple_ports_conjunct_correct: "simple_match_port p1 pkt \ simple_match_port p2 pkt \ simple_match_port (simpl_ports_conjunct p1 p2) pkt" apply(cases p1, cases p2, simp) by blast lemma simple_match_port_code[code] :"simple_match_port (s,e) p_p = (s \ p_p \ p_p \ e)" by simp - lemma simple_match_port_UNIV: "{p. simple_match_port (s,e) p} = UNIV \ (s = 0 \ e = max_word)" + lemma simple_match_port_UNIV: "{p. simple_match_port (s,e) p} = UNIV \ (s = 0 \ e = - 1)" apply(simp) apply(rule) apply(case_tac "s = 0") using antisym_conv apply blast using word_le_0_iff apply blast using word_zero_le by blast subsection\Simple IPs\ lemma simple_match_ip_conjunct: fixes ip1 :: "'i::len word \ nat" shows "simple_match_ip ip1 p_ip \ simple_match_ip ip2 p_ip \ (case ipcidr_conjunct ip1 ip2 of None \ False | Some ipx \ simple_match_ip ipx p_ip)" proof - { fix b1 m1 b2 m2 have "simple_match_ip (b1, m1) p_ip \ simple_match_ip (b2, m2) p_ip \ p_ip \ ipset_from_cidr b1 m1 \ ipset_from_cidr b2 m2" by simp also have "\ \ p_ip \ (case ipcidr_conjunct (b1, m1) (b2, m2) of None \ {} | Some (bx, mx) \ ipset_from_cidr bx mx)" using ipcidr_conjunct_correct by blast also have "\ \ (case ipcidr_conjunct (b1, m1) (b2, m2) of None \ False | Some ipx \ simple_match_ip ipx p_ip)" by(simp split: option.split) finally have "simple_match_ip (b1, m1) p_ip \ simple_match_ip (b2, m2) p_ip \ (case ipcidr_conjunct (b1, m1) (b2, m2) of None \ False | Some ipx \ simple_match_ip ipx p_ip)" . } thus ?thesis by(cases ip1, cases ip2, simp) qed declare simple_matches.simps[simp del] subsection\Merging Simple Matches\ text\@{typ "'i::len simple_match"} \\\ @{typ "'i::len simple_match"}\ fun simple_match_and :: "'i::len simple_match \ 'i simple_match \ 'i simple_match option" where "simple_match_and \iiface=iif1, oiface=oif1, src=sip1, dst=dip1, proto=p1, sports=sps1, dports=dps1 \ \iiface=iif2, oiface=oif2, src=sip2, dst=dip2, proto=p2, sports=sps2, dports=dps2 \ = (case ipcidr_conjunct sip1 sip2 of None \ None | Some sip \ (case ipcidr_conjunct dip1 dip2 of None \ None | Some dip \ (case iface_conjunct iif1 iif2 of None \ None | Some iif \ (case iface_conjunct oif1 oif2 of None \ None | Some oif \ (case simple_proto_conjunct p1 p2 of None \ None | Some p \ Some \iiface=iif, oiface=oif, src=sip, dst=dip, proto=p, sports=simpl_ports_conjunct sps1 sps2, dports=simpl_ports_conjunct dps1 dps2 \)))))" lemma simple_match_and_correct: "simple_matches m1 p \ simple_matches m2 p \ (case simple_match_and m1 m2 of None \ False | Some m \ simple_matches m p)" proof - obtain iif1 oif1 sip1 dip1 p1 sps1 dps1 where m1: "m1 = \iiface=iif1, oiface=oif1, src=sip1, dst=dip1, proto=p1, sports=sps1, dports=dps1 \" by(cases m1, blast) obtain iif2 oif2 sip2 dip2 p2 sps2 dps2 where m2: "m2 = \iiface=iif2, oiface=oif2, src=sip2, dst=dip2, proto=p2, sports=sps2, dports=dps2 \" by(cases m2, blast) have sip_None: "ipcidr_conjunct sip1 sip2 = None \ \ simple_match_ip sip1 (p_src p) \ \ simple_match_ip sip2 (p_src p)" using simple_match_ip_conjunct[of sip1 "p_src p" sip2] by simp have dip_None: "ipcidr_conjunct dip1 dip2 = None \ \ simple_match_ip dip1 (p_dst p) \ \ simple_match_ip dip2 (p_dst p)" using simple_match_ip_conjunct[of dip1 "p_dst p" dip2] by simp have sip_Some: "\ip. ipcidr_conjunct sip1 sip2 = Some ip \ simple_match_ip ip (p_src p) \ simple_match_ip sip1 (p_src p) \ simple_match_ip sip2 (p_src p)" using simple_match_ip_conjunct[of sip1 "p_src p" sip2] by simp have dip_Some: "\ip. ipcidr_conjunct dip1 dip2 = Some ip \ simple_match_ip ip (p_dst p) \ simple_match_ip dip1 (p_dst p) \ simple_match_ip dip2 (p_dst p)" using simple_match_ip_conjunct[of dip1 "p_dst p" dip2] by simp have iiface_None: "iface_conjunct iif1 iif2 = None \ \ match_iface iif1 (p_iiface p) \ \ match_iface iif2 (p_iiface p)" using iface_conjunct[of iif1 "(p_iiface p)" iif2] by simp have oiface_None: "iface_conjunct oif1 oif2 = None \ \ match_iface oif1 (p_oiface p) \ \ match_iface oif2 (p_oiface p)" using iface_conjunct[of oif1 "(p_oiface p)" oif2] by simp have iiface_Some: "\iface. iface_conjunct iif1 iif2 = Some iface \ match_iface iface (p_iiface p) \ match_iface iif1 (p_iiface p) \ match_iface iif2 (p_iiface p)" using iface_conjunct[of iif1 "(p_iiface p)" iif2] by simp have oiface_Some: "\iface. iface_conjunct oif1 oif2 = Some iface \ match_iface iface (p_oiface p) \ match_iface oif1 (p_oiface p) \ match_iface oif2 (p_oiface p)" using iface_conjunct[of oif1 "(p_oiface p)" oif2] by simp have proto_None: "simple_proto_conjunct p1 p2 = None \ \ match_proto p1 (p_proto p) \ \ match_proto p2 (p_proto p)" using simple_proto_conjunct_correct[of p1 "(p_proto p)" p2] by simp have proto_Some: "\proto. simple_proto_conjunct p1 p2 = Some proto \ match_proto proto (p_proto p) \ match_proto p1 (p_proto p) \ match_proto p2 (p_proto p)" using simple_proto_conjunct_correct[of p1 "(p_proto p)" p2] by simp have case_Some: "\m. Some m = simple_match_and m1 m2 \ (simple_matches m1 p \ simple_matches m2 p) \ simple_matches m p" apply(simp add: m1 m2 simple_matches.simps split: option.split_asm) using simple_ports_conjunct_correct by(blast dest: sip_Some dip_Some iiface_Some oiface_Some proto_Some) have case_None: "simple_match_and m1 m2 = None \ \ (simple_matches m1 p \ simple_matches m2 p)" apply(simp add: m1 m2 simple_matches.simps split: option.split_asm) apply(blast dest: sip_None dip_None iiface_None oiface_None proto_None)+ done from case_Some case_None show ?thesis by(cases "simple_match_and m1 m2") simp_all qed lemma simple_match_and_SomeD: "simple_match_and m1 m2 = Some m \ simple_matches m p \ (simple_matches m1 p \ simple_matches m2 p)" by(simp add: simple_match_and_correct) lemma simple_match_and_NoneD: "simple_match_and m1 m2 = None \ \(simple_matches m1 p \ simple_matches m2 p)" by(simp add: simple_match_and_correct) lemma simple_matches_andD: "simple_matches m1 p \ simple_matches m2 p \ \m. simple_match_and m1 m2 = Some m \ simple_matches m p" by (meson option.exhaust_sel simple_match_and_NoneD simple_match_and_SomeD) subsection\Further Properties of a Simple Firewall\ fun has_default_policy :: "'i::len simple_rule list \ bool" where "has_default_policy [] = False" | "has_default_policy [(SimpleRule m _)] = (m = simple_match_any)" | "has_default_policy (_#rs) = has_default_policy rs" lemma has_default_policy: "has_default_policy rs \ simple_fw rs p = Decision FinalAllow \ simple_fw rs p = Decision FinalDeny" proof(induction rs rule: has_default_policy.induct) case 1 thus ?case by (simp) next case (2 m a) thus ?case by(cases a) (simp_all add: simple_match_any) next case (3 r1 r2 rs) from 3 obtain a m where "r1 = SimpleRule m a" by (cases r1) simp with 3 show ?case by (cases a) simp_all qed lemma has_default_policy_fst: "has_default_policy rs \ has_default_policy (r#rs)" apply(cases r, rename_tac m a, simp) apply(cases rs) by(simp_all) text\We can stop after a default rule (a rule which matches anything) is observed.\ fun cut_off_after_match_any :: "'i::len simple_rule list \ 'i simple_rule list" where "cut_off_after_match_any [] = []" | "cut_off_after_match_any (SimpleRule m a # rs) = (if m = simple_match_any then [SimpleRule m a] else SimpleRule m a # cut_off_after_match_any rs)" lemma cut_off_after_match_any: "simple_fw (cut_off_after_match_any rs) p = simple_fw rs p" apply(induction rs p rule: simple_fw.induct) by(simp add: simple_match_any)+ lemma simple_fw_not_matches_removeAll: "\ simple_matches m p \ simple_fw (removeAll (SimpleRule m a) rs) p = simple_fw rs p" apply(induction rs p rule: simple_fw.induct) apply(simp) apply(simp_all) apply blast+ done subsection\Reality check: Validity of Simple Matches\ text\While it is possible to construct a \simple_fw\ expression that only matches a source or destination port, such a match is not meaningful, as the presence of the port information is dependent on the protocol. Thus, a match for a port should always include the match for a protocol. Additionally, prefixes should be zero on bits beyond the prefix length. \ definition "valid_prefix_fw m = valid_prefix (uncurry PrefixMatch m)" lemma ipcidr_conjunct_valid: "\valid_prefix_fw p1; valid_prefix_fw p2; ipcidr_conjunct p1 p2 = Some p\ \ valid_prefix_fw p" unfolding valid_prefix_fw_def by(cases p; cases p1; cases p2) (simp add: ipcidr_conjunct.simps split: if_splits) definition simple_match_valid :: "('i::len, 'a) simple_match_scheme \ bool" where "simple_match_valid m \ ({p. simple_match_port (sports m) p} \ UNIV \ {p. simple_match_port (dports m) p} \ UNIV \ proto m \ Proto `{TCP, UDP, L4_Protocol.SCTP}) \ valid_prefix_fw (src m) \ valid_prefix_fw (dst m)" lemma simple_match_valid_alt[code_unfold]: "simple_match_valid = (\ m. - (let c = (\(s,e). (s \ 0 \ e \ max_word)) in ( + (let c = (\(s,e). (s \ 0 \ e \ - 1)) in ( if c (sports m) \ c (dports m) then proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP else True)) \ valid_prefix_fw (src m) \ valid_prefix_fw (dst m))" proof - - have simple_match_valid_alt_hlp1: "{p. simple_match_port x p} \ UNIV \ (case x of (s,e) \ s \ 0 \ e \ max_word)" + have simple_match_valid_alt_hlp1: "{p. simple_match_port x p} \ UNIV \ (case x of (s,e) \ s \ 0 \ e \ - 1)" for x using simple_match_port_UNIV by auto have simple_match_valid_alt_hlp2: "{p. simple_match_port x p} \ {} \ (case x of (s,e) \ s \ e)" for x by auto show ?thesis unfolding fun_eq_iff unfolding simple_match_valid_def Let_def unfolding simple_match_valid_alt_hlp1 simple_match_valid_alt_hlp2 apply(clarify, rename_tac m, case_tac "sports m"; case_tac "dports m"; case_tac "proto m") by auto qed text\Example:\ context begin private definition "example_simple_match1 \ \iiface = Iface ''+'', oiface = Iface ''+'', src = (0::32 word, 0), dst = (0, 0), proto = Proto TCP, sports = (0, 1024), dports = (0, 1024)\" lemma "simple_fw [SimpleRule example_simple_match1 Drop] \p_iiface = '''', p_oiface = '''', p_src = (1::32 word), p_dst = 2, p_proto = TCP, p_sport = 8, p_dport = 9, p_tcp_flags = {}, p_payload = ''''\ = Decision FinalDeny" by eval private definition "example_simple_match2 \ example_simple_match1\ proto := ProtoAny \" text\Thus, \example_simple_match1\ is valid, but if we set its protocol match to any, it isn't anymore\ private lemma "simple_match_valid example_simple_match1" by eval private lemma "\ simple_match_valid example_simple_match2" by eval end lemma simple_match_and_valid: fixes m1 :: "'i::len simple_match" assumes mv: "simple_match_valid m1" "simple_match_valid m2" assumes mj: "simple_match_and m1 m2 = Some m" shows "simple_match_valid m" proof - have simpl_ports_conjunct_not_UNIV: "Collect (simple_match_port x) \ UNIV \ x = simpl_ports_conjunct p1 p2 \ Collect (simple_match_port p1) \ UNIV \ Collect (simple_match_port p2) \ UNIV" for x p1 p2 by (metis Collect_cong mem_Collect_eq simple_ports_conjunct_correct) (* prefix validity. That's simple. *) have "valid_prefix_fw (src m1)" "valid_prefix_fw (src m2)" "valid_prefix_fw (dst m1)" "valid_prefix_fw (dst m2)" using mv unfolding simple_match_valid_alt by simp_all moreover have "ipcidr_conjunct (src m1) (src m2) = Some (src m)" "ipcidr_conjunct (dst m1) (dst m2) = Some (dst m)" using mj by(cases m1; cases m2; cases m; simp split: option.splits)+ ultimately have pv: "valid_prefix_fw (src m)" "valid_prefix_fw (dst m)" using ipcidr_conjunct_valid by blast+ (* now for the source ports\ *) define nmu where "nmu ps \ {p. simple_match_port ps p} \ UNIV" for ps have "simpl_ports_conjunct (sports m1) (sports m2) = (sports m)" (is "?ph1 sports") using mj by(cases m1; cases m2; cases m; simp split: option.splits) hence sp: "nmu (sports m) \ nmu (sports m1) \ nmu (sports m2)" (is "?ph2 sports") unfolding nmu_def using simpl_ports_conjunct_not_UNIV by metis (* dst ports: mutatis mutandem *) have "?ph1 dports" using mj by(cases m1; cases m2; cases m; simp split: option.splits) hence dp: "?ph2 dports" unfolding nmu_def using simpl_ports_conjunct_not_UNIV by metis (* And an argument for the protocol. *) define php where "php mr \ proto mr \ Proto ` {TCP, UDP, L4_Protocol.SCTP}" for mr :: "'i simple_match" have pcj: "simple_proto_conjunct (proto m1) (proto m2) = Some (proto m)" using mj by(cases m1; cases m2; cases m; simp split: option.splits) hence p: "php m1 \ php m" "php m2 \ php m" using conjunctProtoD conjunctProtoD2 pcj unfolding php_def by auto (* Since I'm trying to trick the simplifier with these defs, I need these as explicit statements: *) have "\mx. simple_match_valid mx \ nmu (sports mx) \ nmu (dports mx) \ php mx" unfolding nmu_def php_def simple_match_valid_def by blast hence mps: "nmu (sports m1) \ php m1" "nmu (dports m1) \ php m1" "nmu (sports m2) \ php m2" "nmu (dports m2) \ php m2" using mv by blast+ have pa: "nmu (sports m) \ nmu (dports m) \ php m" (* apply(intro impI) apply(elim disjE) apply(drule sp[THEN mp]) apply(elim disjE) apply(drule mps) apply(elim p; fail) *) using sp dp mps p by fast show ?thesis unfolding simple_match_valid_def using pv pa[unfolded nmu_def php_def] by blast qed definition "simple_fw_valid \ list_all (simple_match_valid \ match_sel)" text\The simple firewall does not care about tcp flags, payload, or any other packet extensions.\ lemma simple_matches_extended_packet: "simple_matches m \p_iiface = iifce, oiface = oifce, p_src = s, dst = d, p_proto = prot, p_sport = sport, p_dport = dport, tcp_flags = tcp_flags, p_payload = payload1\ \ simple_matches m \p_iiface = iifce, oiface = oifce, p_src = s, p_dst = d, p_proto = prot, p_sport = sport, p_dport = dport, p_tcp_flags = tcp_flags2, p_payload = payload2, \ = aux\ " by(simp add: simple_matches.simps) end diff --git a/thys/Word_Lib/Enumeration_Word.thy b/thys/Word_Lib/Enumeration_Word.thy --- a/thys/Word_Lib/Enumeration_Word.thy +++ b/thys/Word_Lib/Enumeration_Word.thy @@ -1,298 +1,298 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Thomas Sewell *) section "Enumeration Instances for Words" theory Enumeration_Word imports "HOL-Library.Word" More_Word Enumeration Even_More_List begin lemma length_word_enum: "length (enum :: 'a :: len word list) = 2 ^ LENGTH('a)" by (simp add: enum_word_def) lemma fromEnum_unat[simp]: "fromEnum (x :: 'a::len word) = unat x" proof - have "enum ! the_index enum x = x" by (auto intro: nth_the_index) moreover have "the_index enum x < length (enum::'a::len word list)" by (auto intro: the_index_bounded) moreover { fix y assume "of_nat y = x" moreover assume "y < 2 ^ LENGTH('a)" ultimately have "y = unat x" using of_nat_inverse by fastforce } ultimately show ?thesis by (simp add: fromEnum_def enum_word_def) qed lemma toEnum_of_nat[simp]: "n < 2 ^ LENGTH('a) \ (toEnum n :: 'a :: len word) = of_nat n" by (simp add: toEnum_def length_word_enum enum_word_def) instantiation word :: (len) enumeration_both begin definition enum_alt_word_def: "enum_alt \ alt_from_ord (enum :: ('a :: len) word list)" instance by (intro_classes, simp add: enum_alt_word_def) end definition upto_enum_step :: "('a :: len) word \ 'a word \ 'a word \ 'a word list" ("[_ , _ .e. _]") where "upto_enum_step a b c \ if c < a then [] else map (\x. a + x * (b - a)) [0 .e. (c - a) div (b - a)]" (* in the wraparound case, bad things happen. *) lemma maxBound_word: "(maxBound::'a::len word) = -1" by (simp add: maxBound_def enum_word_def last_map of_nat_diff) lemma minBound_word: "(minBound::'a::len word) = 0" by (simp add: minBound_def enum_word_def upt_conv_Cons) lemma maxBound_max_word: - "(maxBound::'a::len word) = max_word" + "(maxBound::'a::len word) = - 1" by (fact maxBound_word) lemma leq_maxBound [simp]: "(x::'a::len word) \ maxBound" by (simp add: maxBound_max_word) lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < LENGTH('a)" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "x < y \ [x::'a::len word .e. y] = x # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons) apply simp_all apply unat_arith apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms) apply simp_all apply unat_arith done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply simp_all apply (metis le_unat_uoi nat_less_le toEnum_of_nat unsigned_less word_le_nat_alt) apply (metis le_unat_uoi less_or_eq_imp_le toEnum_of_nat unsigned_less word_le_nat_alt) apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply transfer apply auto done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv le_m1_iff_lt p2_gt_0 by auto qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto lemma length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_cases: fixes a :: "'a::len word" shows "length [a .e. b] = (if a \ b then Suc (unat b) - unat a else 0)" apply (case_tac "a \ b") apply (clarsimp) apply (clarsimp simp: upto_enum_def) apply unat_arith done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done lemma enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done end diff --git a/thys/Word_Lib/Guide.thy b/thys/Word_Lib/Guide.thy --- a/thys/Word_Lib/Guide.thy +++ b/thys/Word_Lib/Guide.thy @@ -1,373 +1,376 @@ (* * Copyright Florian Haftmann * * SPDX-License-Identifier: BSD-2-Clause *) (*<*) theory Guide imports Word_Lib_Sumo Word_64 Ancient_Numeral begin hide_const (open) Generic_set_bit.set_bit (*>*) section \A short overview over bit operations and word types\ subsection \Basic theories and key ideas\ text \ When formalizing bit operations, it is tempting to represent bit values as explicit lists over a binary type. This however is a bad idea, mainly due to the inherent ambiguities in representation concerning repeating leading bits. Hence this approach avoids such explicit lists altogether following an algebraic path: \<^item> Bit values are represented by numeric types: idealized unbounded bit values can be represented by type \<^typ>\int\, bounded bit values by quotient types over \<^typ>\int\, aka \<^typ>\'a word\. \<^item> (A special case are idealized unbounded bit values ending in @{term [source] 0} which can be represented by type \<^typ>\nat\ but only support a restricted set of operations). The most fundamental ideas are developed in theory \<^theory>\HOL.Parity\ (which is part of \<^theory>\Main\): \<^item> Multiplication by \<^term>\2 :: int\ is a bit shift to the left and \<^item> Division by \<^term>\2 :: int\ is a bit shift to the right. \<^item> Concerning bounded bit values, iterated shifts to the left may result in eliminating all bits by shifting them all beyond the boundary. The property \<^prop>\(2 :: int) ^ n \ 0\ represents that \<^term>\n\ is \<^emph>\not\ beyond that boundary. \<^item> The projection on a single bit is then @{thm [mode=iff] bit_iff_odd [where ?'a = int, no_vars]}. \<^item> This leads to the most fundamental properties of bit values: \<^item> Equality rule: @{thm [display, mode=iff] bit_eq_iff [where ?'a = int, no_vars]} \<^item> Induction rule: @{thm [display, mode=iff] bits_induct [where ?'a = int, no_vars]} \<^item> Characteristic properties \<^prop>\bit (f x) n \ P x n\ are available in fact collection \<^text>\bit_simps\. On top of this, the following generic operations are provided after import of theory \<^theory>\HOL-Library.Bit_Operations\: \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ \<^item> Bit mask upto bit \<^term>\n\: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]} \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} \<^item> Bitwise negation: @{thm [mode=iff] bit_not_iff [where ?'a = int, no_vars]} \<^item> Bitwise conjunction: @{thm [mode=iff] bit_and_iff [where ?'a = int, no_vars]} \<^item> Bitwise disjunction: @{thm [mode=iff] bit_or_iff [where ?'a = int, no_vars]} \<^item> Bitwise exclusive disjunction: @{thm [mode=iff] bit_xor_iff [where ?'a = int, no_vars]} \<^item> Setting a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} \<^item> Unsetting a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} \<^item> Flipping a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm [display] signed_take_bit_def [where ?'a = int, no_vars]} \<^item> (Bounded) conversion from and to a list of bits: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} Bit concatenation on \<^typ>\int\ as given by @{thm [display] concat_bit_def [no_vars]} appears quite technical but is the logical foundation for the quite natural bit concatenation on \<^typ>\'a word\ (see below). Proper word types are introduced in theory \<^theory>\HOL-Library.Word\, with the following specific operations: \<^item> Standard arithmetic: @{term \(+) :: 'a::len word \ 'a word \ 'a word\}, @{term \uminus :: 'a::len word \ 'a word\}, @{term \(-) :: 'a::len word \ 'a word \ 'a word\}, @{term \(*) :: 'a::len word \ 'a word \ 'a word\}, @{term \0 :: 'a::len word\}, @{term \1 :: 'a::len word\}, numerals etc. \<^item> Standard bit operations: see above. \<^item> Conversion with unsigned interpretation of words: \<^item> @{term [source] \unsigned :: 'a::len word \ 'b::semiring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \unat :: 'a::len word \ nat\} \<^item> @{term [source] \uint :: 'a::len word \ int\} \<^item> @{term [source] \ucast :: 'a::len word \ 'b::len word\} \<^item> Conversion with signed interpretation of words: \<^item> @{term [source] \signed :: 'a::len word \ 'b::ring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \sint :: 'a::len word \ int\} \<^item> @{term [source] \scast :: 'a::len word \ 'b::len word\} \<^item> Operations with unsigned interpretation of words: \<^item> @{thm [mode=iff] word_le_nat_alt [no_vars]} \<^item> @{thm [mode=iff] word_less_nat_alt [no_vars]} \<^item> @{thm unat_div_distrib [no_vars]} \<^item> @{thm unat_drop_bit_eq [no_vars]} \<^item> @{thm unat_mod_distrib [no_vars]} \<^item> @{thm [mode=iff] udvd_iff_dvd [no_vars]} \<^item> Operations with signed interpretation of words: \<^item> @{thm [mode=iff] word_sle_eq [no_vars]} \<^item> @{thm [mode=iff] word_sless_alt [no_vars]} \<^item> @{thm sint_signed_drop_bit_eq [no_vars]} \<^item> Rotation and reversal: \<^item> @{term [source] \word_rotl :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_rotr :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_roti :: int \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_reverse :: 'a::len word \ 'a word\} \<^item> Concatenation: @{term [source, display] \word_cat :: 'a::len word \ 'b::len word \ 'c::len word\} For proofs about words the following default strategies are applicable: \<^item> Using bit extensionality (facts \<^text>\bit_eq_iff\, \<^text>\bit_eqI\; fact collection \<^text>\bit_simps\). \<^item> Using the @{method transfer} method. \ subsection \More library theories\ text \ Note: currently, the theories listed here are hardly separate entities since they import each other in various ways. Always inspect them to understand what you pull in if you want to import one. \<^descr>[Syntax] \<^descr>[\<^theory>\Word_Lib.Hex_Words\] Printing word numerals as hexadecimal numerals. \<^descr>[\<^theory>\Word_Lib.Type_Syntax\] Pretty type-sensitive syntax for cast operations. \<^descr>[\<^theory>\Word_Lib.Word_Syntax\] Specific ASCII syntax for prominent bit operations on word. \<^descr>[Proof tools] \<^descr>[\<^theory>\Word_Lib.Norm_Words\] Rewriting word numerals to normal forms. \<^descr>[\<^theory>\Word_Lib.Bitwise\] Method @{method word_bitwise} decomposes word equalities and inequalities into bit propositions. \<^descr>[\<^theory>\Word_Lib.Word_EqI\] Method @{method word_eqI_solve} decomposes word equalities and inequalities into bit propositions. \<^descr>[Operations] \<^descr>[\<^theory>\Word_Lib.Signed_Division_Word\] Signed division on word: \<^item> @{term [source] \(sdiv) :: 'a::len word \ 'a word \ 'a word\} \<^item> @{term [source] \(smod) :: 'a::len word \ 'a word \ 'a word\} \<^descr>[\<^theory>\Word_Lib.Aligned\] \ \<^item> @{thm [mode=iff] is_aligned_iff_udvd [no_vars]} \<^descr>[\<^theory>\Word_Lib.Least_significant_bit\] The least significant bit as an alias: @{thm [mode=iff] lsb_odd [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Most_significant_bit\] The most significant bit: \<^item> @{thm [mode=iff] msb_int_def [of k]} \<^item> @{thm [mode=iff] word_msb_sint [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_sless_0 [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_bit [no_vars]} \<^descr>[\<^theory>\Word_Lib.Traditional_Infix_Syntax\] Clones of existing operations decorated with traditional syntax: \<^item> @{thm test_bit_eq_bit [no_vars]} \<^item> @{thm shiftl_eq_push_bit [no_vars]} \<^item> @{thm shiftr_eq_drop_bit [no_vars]} \<^item> @{thm sshiftr_eq [no_vars]} \<^descr>[\<^theory>\Word_Lib.Next_and_Prev\] \ \<^item> @{thm word_next_unfold [no_vars]} \<^item> @{thm word_prev_unfold [no_vars]} \<^descr>[\<^theory>\Word_Lib.Enumeration_Word\] More on explicit enumeration of word types. \<^descr>[\<^theory>\Word_Lib.More_Word_Operations\] Even more operations on word. \<^descr>[Types] \<^descr>[\<^theory>\Word_Lib.Signed_Words\] Formal tagging of word types with a \<^text>\signed\ marker. \<^descr>[Lemmas] \<^descr>[\<^theory>\Word_Lib.More_Word\] More lemmas on words. \<^descr>[\<^theory>\Word_Lib.Word_Lemmas\] More lemmas on words, covering many other theories mentioned here. \<^descr>[Words of popular lengths]. \<^descr>[\<^theory>\Word_Lib.Word_8\] for 8-bit words. \<^descr>[\<^theory>\Word_Lib.Word_16\] for 16-bit words. \<^descr>[\<^theory>\Word_Lib.Word_32\] for 32-bit words. \<^descr>[\<^theory>\Word_Lib.Word_64\] for 64-bit words. This theory is not part of \<^text>\Word_Lib_Sumo\, because it shadows names from \<^theory>\Word_Lib.Word_32\. They can be used together, but then require to use qualified names in applications. \ subsection \More library sessions\ text \ \<^descr>[\<^text>\Native_Word\] Makes machine words and machine arithmetic available for code generation. It provides a common abstraction that hides the differences between the different target languages. The code generator maps these operations to the APIs of the target languages. \ subsection \Legacy theories\ text \ The following theories contain material which has been factored out since it is not recommended to use it in new applications, mostly because matters can be expressed succinctly using already existing operations. This section gives some indication how to migrate away from those theories. However theorem coverage may still be terse in some cases. \<^descr>[\<^theory>\Word_Lib.Word_Lib_Sumo\] An entry point importing any relevant theory in that session. Intended for backward compatibility: start importing this theory when migrating applications to Isabelle2021, and later sort out what you really need. You may need to include \<^theory>\Word_Lib.Word_64\ separately. \<^descr>[\<^theory>\Word_Lib.Generic_set_bit\] Kind of an alias: @{thm set_bit_eq [no_vars]} \<^descr>[\<^theory>\Word_Lib.Typedef_Morphisms\] A low-level extension to HOL typedef providing conversions along type morphisms. The @{method transfer} method seems to be sufficient for most applications though. \<^descr>[\<^theory>\Word_Lib.Bit_Comprehension\] Comprehension syntax for bit values over predicates \<^typ>\nat \ bool\. For \<^typ>\'a::len word\, straightforward alternatives exist; difficult to handle for \<^typ>\int\. \<^descr>[\<^theory>\Word_Lib.Reversed_Bit_Lists\] Representation of bit values as explicit list in \<^emph>\reversed\ order. This should rarely be necessary: the \<^const>\bit\ projection should be sufficient in most cases. In case explicit lists are needed, existing operations can be used: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Many_More\] Collection of operations and theorems which are kept for backward compatibility and not used in other theories in session \<^text>\Word_Lib\. They are used in applications of \<^text>\Word_Lib\, but should be migrated to there. \ section \Changelog\ text \ \<^descr>[Changes since AFP 2021] ~ - \<^item> Theory \<^theory>\Word_Lib.Ancient_Numeral\ is not part of \<^theory>\Word_Lib.Word_Lib_Sumo\ + \<^item> Theory \<^theory>\Word_Lib.Ancient_Numeral\ is no part of \<^theory>\Word_Lib.Word_Lib_Sumo\ any longer. + + \<^item> Abbreviation \<^abbrev>\max_word\ moved from distribution into theory + \<^theory>\Word_Lib.Legacy_Aliases\. \ (*<*) end (*>*) diff --git a/thys/Word_Lib/Legacy_Aliases.thy b/thys/Word_Lib/Legacy_Aliases.thy --- a/thys/Word_Lib/Legacy_Aliases.thy +++ b/thys/Word_Lib/Legacy_Aliases.thy @@ -1,24 +1,27 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Legacy aliases\ theory Legacy_Aliases imports "HOL-Library.Word" begin +abbreviation (input) max_word :: \'a::len word\ + where "max_word \ - 1" + definition complement :: "'a :: len word \ 'a word" where "complement x \ NOT x" lemma complement_mask: "complement (2 ^ n - 1) = NOT (mask n)" unfolding complement_def mask_eq_decr_exp by simp lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] end diff --git a/thys/Word_Lib/More_Word.thy b/thys/Word_Lib/More_Word.thy --- a/thys/Word_Lib/More_Word.thy +++ b/thys/Word_Lib/More_Word.thy @@ -1,1807 +1,1807 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Lemmas on words\ theory More_Word imports "HOL-Library.Word" More_Arithmetic More_Divides begin lemma unat_power_lower [simp]: "unat ((2::'a::len word) ^ n) = 2 ^ n" if "n < LENGTH('a::len)" using that by transfer simp lemma unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" by (fact unat_power_lower) lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by transfer simp lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ unat n < 2 * unat (m :: 'a :: len word)" apply (simp only: word_arith_nat_defs word_le_nat_alt word_of_nat_eq_iff flip: nat_div_eq_Suc_0_iff) apply (simp flip: unat_div unsigned_take_bit_eq) done lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" apply (induct x) apply simp apply (simp add: shiftl1_2t) done lemma AND_twice [simp]: "(w AND m) AND m = w AND m" by (fact and.right_idem) lemma word_combine_masks: "w AND m = z \ w AND m' = z' \ w AND (m OR m') = (z OR z')" for w m m' z z' :: \'a::len word\ by (simp add: bit.conj_disj_distrib) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" by (simp add : word_gt_0 not_le) lemma uint_2p_alt: \n < LENGTH('a::len) \ uint ((2::'a::len word) ^ n) = 2 ^ n\ using p2_gt_0 [of n, where ?'a = 'a] by (simp add: uint_2p) lemma p2_eq_0: \(2::'a::len word) ^ n = 0 \ LENGTH('a::len) \ n\ by (fact exp_eq_zero_iff) lemma p2len: \(2 :: 'a word) ^ LENGTH('a::len) = 0\ by simp lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps simp flip: push_bit_eq_mult drop_bit_eq_div) lemma neg_mask_is_div': "n < size w \ w AND NOT (mask n) = ((w div (2 ^ n)) * (2 ^ n))" for w :: \'a::len word\ by (rule neg_mask_is_div) lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size simp flip: push_bit_eq_mult drop_bit_eq_div) lemma and_mask_arith': "0 < n \ w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule and_mask_arith) lemma mask_2pm1: "mask n = 2 ^ n - (1 :: 'a::len word)" by (fact mask_eq_decr_exp) lemma add_mask_fold: "x + 2 ^ n - 1 = x + mask n" for x :: \'a::len word\ by (simp add: mask_eq_decr_exp) lemma word_and_mask_le_2pm1: "w AND mask n \ 2 ^ n - 1" for w :: \'a::len word\ by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u AND mask n = 0 \ v < 2^n \ u AND v = 0" for u v :: \'a::len word\ apply (drule less_mask_eq) apply (simp flip: take_bit_eq_mask) apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps) done lemma le_shiftr1: \shiftr1 u \ shiftr1 v\ if \u \ v\ using that proof transfer fix k l :: int assume \take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ then have \take_bit LENGTH('a) (drop_bit 1 (take_bit LENGTH('a) k)) \ take_bit LENGTH('a) (drop_bit 1 (take_bit LENGTH('a) l))\ apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done then show \take_bit LENGTH('a) (take_bit LENGTH('a) k div 2) \ take_bit LENGTH('a) (take_bit LENGTH('a) l div 2)\ by (simp add: drop_bit_eq_div) qed lemma and_mask_eq_iff_le_mask: \w AND mask n = w \ w \ mask n\ for w :: \'a::len word\ apply (simp flip: take_bit_eq_mask) apply (cases \n \ LENGTH('a)\; transfer) apply (simp_all add: not_le min_def) apply (simp_all add: mask_eq_exp_minus_1) apply auto apply (metis take_bit_int_less_exp) apply (metis min_def nat_less_le take_bit_int_eq_self_iff take_bit_take_bit) done lemma less_eq_mask_iff_take_bit_eq_self: \w \ mask n \ take_bit n w = w\ for w :: \'a::len word\ by (simp add: and_mask_eq_iff_le_mask take_bit_eq_mask) lemma NOT_eq: "NOT (x :: 'a :: len word) = - x - 1" apply (cut_tac x = "x" in word_add_not) apply (drule add.commute [THEN trans]) apply (drule eq_diff_eq [THEN iffD2]) by simp lemma NOT_mask: "NOT (mask n :: 'a::len word) = - (2 ^ n)" by (simp add : NOT_eq mask_2pm1) lemma le_m1_iff_lt: "(x > (0 :: 'a :: len word)) = ((y \ x - 1) = (y < x))" by uint_arith lemma gt0_iff_gem1: \0 < x \ x - 1 < x\ for x :: \'a::len word\ by (metis add.right_neutral diff_add_cancel less_irrefl measure_unat unat_arith_simps(2) word_neq_0_conv word_sub_less_iff) lemma power_2_ge_iff: \2 ^ n - (1 :: 'a::len word) < 2 ^ n \ n < LENGTH('a)\ using gt0_iff_gem1 p2_gt_0 by blast lemma le_mask_iff_lt_2n: "n < len_of TYPE ('a) = (((w :: 'a :: len word) \ mask n) = (w < 2 ^ n))" unfolding mask_2pm1 by (rule trans [OF p2_gt_0 [THEN sym] le_m1_iff_lt]) lemma mask_lt_2pn: \n < LENGTH('a) \ mask n < (2 :: 'a::len word) ^ n\ by (simp add: mask_eq_exp_minus_1 power_2_ge_iff) lemma word_unat_power: "(2 :: 'a :: len word) ^ n = of_nat (2 ^ n)" by simp lemma of_nat_mono_maybe: assumes xlt: "x < 2 ^ len_of TYPE ('a)" shows "y < x \ of_nat y < (of_nat x :: 'a :: len word)" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (erule order_less_trans [OF _ xlt]) apply (subst mod_less [OF xlt]) apply assumption done lemma word_and_max_word: fixes a::"'a::len word" - shows "x = max_word \ a AND x = a" + shows "x = - 1 \ a AND x = a" by simp lemma word_and_full_mask_simp: \x AND mask LENGTH('a) = x\ for x :: \'a::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ (0 :: 'a word)\ then have \n < LENGTH('a)\ by simp then show \bit (x AND Bit_Operations.mask LENGTH('a)) n \ bit x n\ by (simp add: bit_and_iff bit_mask_iff) qed lemma of_int_uint: "of_int (uint x) = x" by (fact word_of_int_uint) corollary word_plus_and_or_coroll: "x AND y = 0 \ x + y = x OR y" for x y :: \'a::len word\ using word_plus_and_or[where x=x and y=y] by simp corollary word_plus_and_or_coroll2: "(x AND w) + (x AND NOT w) = x" for x w :: \'a::len word\ apply (subst disjunctive_add) apply (simp add: bit_simps) apply (simp flip: bit.conj_disj_distrib) done lemma nat_mask_eq: \nat (mask n) = mask n\ by (simp add: nat_eq_iff of_nat_mask_eq) lemma unat_mask_eq: \unat (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer (simp add: nat_mask_eq) lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by transfer (simp add: min_def split: if_splits) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" by (metis diff_minus_eq_add less_imp_le sub_wrap_lt) lemma ucast_ucast_eq: "\ ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" for x :: "'a::len word" and y :: "'b::len word" apply transfer apply (cases \LENGTH('c) = LENGTH('a)\) apply (auto simp add: min_def) done lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0"; simp) apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less; simp) done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_eq_unatI: \v = w\ if \unat v = unat w\ using that by transfer (simp add: nat_eq_iff) lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_eq_unatI) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma enum_word_nth_eq: \(Enum.enum :: 'a::len word list) ! n = word_of_nat n\ if \n < 2 ^ LENGTH('a)\ for n using that by (simp add: enum_word_def) lemma length_enum_word_eq: \length (Enum.enum :: 'a::len word list) = 2 ^ LENGTH('a)\ by (simp add: enum_word_def) lemma unat_lt2p [iff]: \unat x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply simp_all done lemma word_sub_1_le: "x \ 0 \ x - 1 \ (x :: ('a :: len) word)" apply (subst no_ulen_sub) apply simp apply (cases "uint x = 0") apply (simp add: uint_0_iff) apply (insert uint_ge_0[where x=x]) apply arith done lemma push_bit_word_eq_nonzero: \push_bit n w \ 0\ if \w < 2 ^ m\ \m + n < LENGTH('a)\ \w \ 0\ for w :: \'a::len word\ using that apply (simp only: word_neq_0_conv word_less_nat_alt mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (metis add_diff_cancel_right' gr0I gr_implies_not0 less_or_eq_imp_le min_def push_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_push_bit take_bit_take_bit unsigned_push_bit_eq) done lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof (cases \sz < LENGTH('a)\) case True with assms show ?thesis by (simp add: unat_word_ariths take_bit_eq_mod mod_simps) (simp add: take_bit_nat_eq_self_iff nat_less_power_trans flip: take_bit_eq_mod) next case False with assms show ?thesis by simp qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma word_le_minus_one_leq: "x < y \ x \ y - 1" for x :: "'a :: len word" by transfer (metis le_less_trans less_irrefl take_bit_decr_eq take_bit_nonnegative zle_diff1_eq) lemma word_less_sub_le[simp]: fixes x :: "'a :: len word" assumes nv: "n < LENGTH('a)" shows "(x \ 2 ^ n - 1) = (x < 2 ^ n)" using le_less_trans word_le_minus_one_leq nv power_2_ge_iff by blast lemma unat_of_nat_len: "x < 2 ^ LENGTH('a) \ unat (of_nat x :: 'a::len word) = x" by (simp add: take_bit_nat_eq_self_iff) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by transfer (auto simp add: take_bit_of_nat nat_eq_iff take_bit_nat_eq_self_iff intro: sym) lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma mask_out_sub_mask: "(x AND NOT (mask n)) = x - (x AND (mask n))" for x :: \'a::len word\ by (simp add: field_simps word_plus_and_or_coroll2) lemma subtract_mask: "p - (p AND mask n) = (p AND NOT (mask n))" "p - (p AND NOT (mask n)) = (p AND mask n)" for p :: \'a::len word\ by (simp add: field_simps word_plus_and_or_coroll2)+ lemma take_bit_word_eq_self_iff: \take_bit n w = w \ n \ LENGTH('a) \ w < 2 ^ n\ for w :: \'a::len word\ using take_bit_int_eq_self_iff [of n \take_bit LENGTH('a) (uint w)\] by (transfer fixing: n) auto lemma word_power_increasing: assumes x: "2 ^ x < (2 ^ y::'a::len word)" "x < LENGTH('a::len)" "y < LENGTH('a::len)" shows "x < y" using x using assms by transfer simp lemma mask_twice: "(x AND mask n) AND mask m = x AND mask (min m n)" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask) lemma plus_one_helper[elim!]: "x < n + (1 :: 'a :: len word) \ x \ n" apply (simp add: word_less_nat_alt word_le_nat_alt field_simps) apply (case_tac "1 + n = 0") apply simp_all apply (subst(asm) unatSuc, assumption) apply arith done lemma plus_one_helper2: "\ x \ n; n + 1 \ 0 \ \ x < n + (1 :: 'a :: len word)" by (simp add: word_less_nat_alt word_le_nat_alt field_simps unatSuc) lemma less_x_plus_1: fixes x :: "'a :: len word" shows - "x \ max_word \ (y < (x + 1)) = (y < x \ y = x)" + "x \ - 1 \ (y < (x + 1)) = (y < x \ y = x)" apply (rule iffI) apply (rule disjCI) apply (drule plus_one_helper) apply simp apply (subgoal_tac "x < x + 1") apply (erule disjE, simp_all) apply (rule plus_one_helper2 [OF order_refl]) apply (rule notI, drule max_word_wrap) apply simp done lemma word_Suc_leq: - fixes k::"'a::len word" shows "k \ max_word \ x < k + 1 \ x \ k" + fixes k::"'a::len word" shows "k \ - 1 \ x < k + 1 \ x \ k" using less_x_plus_1 word_le_less_eq by auto lemma word_Suc_le: - fixes k::"'a::len word" shows "x \ max_word \ x + 1 \ k \ x < k" + fixes k::"'a::len word" shows "x \ - 1 \ x + 1 \ k \ x < k" by (meson not_less word_Suc_leq) lemma word_lessThan_Suc_atMost: \{..< k + 1} = {..k}\ if \k \ - 1\ for k :: \'a::len word\ using that by (simp add: lessThan_def atMost_def word_Suc_leq) lemma word_atLeastLessThan_Suc_atLeastAtMost: \{l ..< u + 1} = {l..u}\ if \u \ - 1\ for l :: \'a::len word\ using that by (simp add: atLeastAtMost_def atLeastLessThan_def word_lessThan_Suc_atMost) lemma word_atLeastAtMost_Suc_greaterThanAtMost: \{m<..u} = {m + 1..u}\ if \m \ - 1\ for m :: \'a::len word\ using that by (simp add: greaterThanAtMost_def greaterThan_def atLeastAtMost_def atLeast_def word_Suc_le) lemma word_atLeastLessThan_Suc_atLeastAtMost_union: fixes l::"'a::len word" - assumes "m \ max_word" and "l \ m" and "m \ u" + assumes "m \ - 1" and "l \ m" and "m \ u" shows "{l..m} \ {m+1..u} = {l..u}" proof - from ivl_disj_un_two(8)[OF assms(2) assms(3)] have "{l..u} = {l..m} \ {m<..u}" by blast with assms show ?thesis by(simp add: word_atLeastAtMost_Suc_greaterThanAtMost) qed lemma max_word_less_eq_iff [simp]: \- 1 \ w \ w = - 1\ for w :: \'a::len word\ by (fact word_order.extremum_unique) lemma word_or_zero: "(a OR b = 0) = (a = 0 \ b = 0)" for a b :: \'a::len word\ by (fact or_eq_0_iff) lemma word_2p_mult_inc: assumes x: "2 * 2 ^ n < (2::'a::len word) * 2 ^ m" assumes suc_n: "Suc n < LENGTH('a::len)" shows "2^n < (2::'a::len word)^m" by (smt suc_n le_less_trans lessI nat_less_le nat_mult_less_cancel_disj p2_gt_0 power_Suc power_Suc unat_power_lower word_less_nat_alt x) lemma power_overflow: "n \ LENGTH('a) \ 2 ^ n = (0 :: 'a::len word)" by simp lemmas extra_sle_sless_unfolds [simp] = word_sle_eq[where a=0 and b=1] word_sle_eq[where a=0 and b="numeral n"] word_sle_eq[where a=1 and b=0] word_sle_eq[where a=1 and b="numeral n"] word_sle_eq[where a="numeral n" and b=0] word_sle_eq[where a="numeral n" and b=1] word_sless_alt[where a=0 and b=1] word_sless_alt[where a=0 and b="numeral n"] word_sless_alt[where a=1 and b=0] word_sless_alt[where a=1 and b="numeral n"] word_sless_alt[where a="numeral n" and b=0] word_sless_alt[where a="numeral n" and b=1] for n lemma word_sint_1: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma ucast_of_nat: "is_down (ucast :: 'a :: len word \ 'b :: len word) \ ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)" by transfer simp lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (signed_take_bit (LENGTH('a::len) - Suc 0) (1::int)))" by transfer simp lemma scast_1: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" apply (simp only: flip: mask_eq_exp_minus_1) apply transfer apply (simp add: take_bit_minus_one_eq_mask nat_mask_eq) done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma word_leq_le_minus_one: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma neg_mask_combine: "NOT(mask a) AND NOT(mask b) = NOT(mask (max a b) :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma neg_mask_twice: "x AND NOT(mask n) AND NOT(mask m) = x AND NOT(mask (max n m))" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma multiple_mask_trivia: "n \ m \ (x AND NOT(mask n)) + (x AND mask n AND NOT(mask m)) = x AND NOT(mask m)" for x :: \'a::len word\ apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2) apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice max_absorb2) done lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_eq_decr_exp not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_eq_decr_exp) lemma Suc_2p_unat_mask: "n < LENGTH('a) \ Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)" by (simp add: unat_mask) lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: bit_iff_odd) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" apply transfer apply (rule signed_take_bit_int_eq_self) apply simp_all apply (metis negative_zle numeral_power_eq_of_nat_cancel_iff) done lemma sint_of_nat_le: "\ b < 2 ^ (LENGTH('a) - 1); a \ b \ \ sint (of_nat a :: 'a :: len word) \ sint (of_nat b :: 'a :: len word)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le order_less_le_trans take_bit_nat_eq_self_iff) apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le take_bit_nat_eq_self_iff) apply (simp flip: of_nat_take_bit add: take_bit_nat_eq_self) done lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by transfer simp lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) apply (metis bit_take_bit_iff less_mask_eq not_less take_bit_eq_mask) done lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" by (fact word_le_minus_one_leq) lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: take_bit_eq_mod) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_eq_nat_uint [symmetric]) apply (subst unat_eq_nat_uint [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" by transfer (simp add: take_bit_eq_mod) lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (metis unat_of_nat_len) lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x AND mask (len_of TYPE ('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma ucast_ucast_id: "LENGTH('a) < LENGTH('b) \ ucast (ucast (x::'a::len word)::'b::len word) = x" by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size) lemma unat_ucast: "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))" proof - have \2 ^ LENGTH('a) = nat (2 ^ LENGTH('a))\ by simp moreover have \unat (ucast x :: 'a word) = unat x mod nat (2 ^ LENGTH('a))\ by transfer (simp flip: nat_mod_distrib take_bit_eq_mod) ultimately show ?thesis by (simp only:) qed lemma ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done \ \This weaker version was previously called @{text ucast_less_ucast}. We retain it to support existing proofs.\ lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order] lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (fact bits_div_by_1) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" by (fact word_order.extremum_unique) lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply transfer apply (cases \LENGTH('a)\) apply simp_all apply (metis order_refl take_bit_signed_take_bit take_bit_tightened) done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma unat_eq_1: \unat x = Suc 0 \ x = 1\ by (auto intro!: unsigned_word_eqI [where ?'a = nat]) lemma word_unat_Rep_inject1: \unat x = unat 1 \ x = 1\ by (simp add: unat_eq_1) lemma and_not_mask_twice: "(w AND NOT (mask n)) AND NOT (mask m) = w AND NOT (mask (max m n))" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_less_cases: "x < y \ x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma mask_and_mask: "mask a AND mask b = (mask (min a b) :: 'a::len word)" by (simp flip: take_bit_eq_mask ac_simps) lemma mask_eq_0_eq_x: "(x AND w = 0) = (x AND NOT w = x)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x AND w = x) = (x AND NOT w = 0)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma compl_of_1: "NOT 1 = (-2 :: 'a :: len word)" by (fact not_one) lemma split_word_eq_on_mask: "(x = y) = (x AND m = y AND m \ x AND NOT m = y AND NOT m)" for x y m :: \'a::len word\ apply transfer apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps ac_simps) done lemma word_FF_is_mask: "0xFF = (mask 8 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma word_1FF_is_mask: "0x1FF = (mask 9 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply transfer apply (auto simp add: take_bit_of_nat min_def not_le) apply (metis linorder_not_less min_def take_bit_nat_eq_self take_bit_take_bit) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma word_leq_minus_one_le: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule word_le_minus_one_leq) apply (rule disjCI) apply simp apply simp apply (erule word_leq_minus_one_le) apply fastforce done lemma word_minus_one_le_leq: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma word_div_less: "m < n \ m div n = 0" for m :: "'a :: len word" by (simp add: unat_mono word_arith_nat_defs(6)) lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" using word_sub_le word_sub_mono by fastforce lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by transfer simp lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply rule apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) apply simp done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp_all add: word_less_nat_alt less_le_trans take_bit_eq_mod) done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (cases \2 \ LENGTH('a)\) apply simp_all apply transfer apply auto apply (metis add.commute add.right_neutral even_two_times_div_two mod_div_trivial mod_pos_pos_trivial mult.commute mult_zero_left not_less not_take_bit_negative odd_two_times_div_two_succ) done lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power del: of_nat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w AND mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply auto using of_nat_inj apply blast done lemma mask_AND_NOT_mask: "(w AND NOT (mask n)) AND mask n = 0" for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_simps) lemma AND_NOT_mask_plus_AND_mask_eq: "(w AND NOT (mask n)) + (w AND mask n) = w" for w :: \'a::len word\ apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x AND mask n = y AND mask n" and m2: "x AND NOT (mask n) = y AND NOT (mask n)" shows "x = y" proof - have *: \x = x AND mask n OR x AND NOT (mask n)\ for x :: \'a word\ by (rule bit_word_eqI) (auto simp add: bit_simps) from assms * [of x] * [of y] show ?thesis by simp qed lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x AND 1) = (x AND 1 = 1)" by (simp add: and_one_eq mod_2_eq_odd) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply transfer apply simp apply (subst (2) take_bit_add [symmetric]) apply (subst take_bit_add [symmetric]) apply simp done lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma word_0_sle_from_less: \0 \s x\ if \x < 2 ^ (LENGTH('a) - 1)\ for x :: \'a::len word\ using that apply transfer apply (cases \LENGTH('a)\) apply simp_all apply (metis bit_take_bit_iff min_def nat_less_le not_less_eq take_bit_int_eq_self_iff take_bit_take_bit) done lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" apply transfer apply (subst (asm) take_bit_incr_eq) apply (auto simp add: diff_less_eq) using take_bit_int_less_exp le_less_trans by blast lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp flip: signed_take_bit_eq_iff_take_bit_eq) done lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (fact signed_eq_0_iff) (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: P if \\ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P\ \\ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P\ \\ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P\ proof (cases \LENGTH('a) = 1\) case True then have \a = 0 \ a = 1\ by transfer auto with True that show ?thesis by auto next case False with that show ?thesis by (simp add: less_le Suc_le_eq) qed lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp add: signed_take_bit_int_eq_self) done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply simp done lemma uint_range': \0 \ uint x \ uint x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" by (simp add: signed_take_bit_int_eq_self) lemma of_int_sint: "of_int (sint a) = a" by simp lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply transfer apply (simp add: signed_take_bit_take_bit) done lemma word_less_nowrapI': "(x :: 'a :: len word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = (2 ^ n :: 'a::len word)" by (clarsimp simp: mask_eq_decr_exp) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_eq unat_eq_nat_uint apply transfer apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp only: flip: ucast_nat_def) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply simp apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp only: flip: ucast_nat_def) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply simp apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp_all add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ ucast x \ (ucast y :: 'b word)" for x y :: \'a::len word\ by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma neg_mask_add_mask: "((x:: 'a :: len word) AND NOT (mask n)) + (2 ^ n - 1) = x OR mask n" unfolding mask_2pm1 [symmetric] apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w AND mask n) AND NOT (mask n) = 0" by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff bit_mask_iff) lemma and_and_not[simp]:"(a AND b) AND NOT b = 0" for a b :: \'a::len word\ apply (subst word_bw_assocs(1)) apply clarsimp done lemma ex_mask_1[simp]: "(\x. mask x = (1 :: 'a::len word))" apply (rule_tac x=1 in exI) apply (simp add:mask_eq_decr_exp) done lemma not_switch:"NOT a = x \ a = NOT x" by auto end diff --git a/thys/Word_Lib/Reversed_Bit_Lists.thy b/thys/Word_Lib/Reversed_Bit_Lists.thy --- a/thys/Word_Lib/Reversed_Bit_Lists.thy +++ b/thys/Word_Lib/Reversed_Bit_Lists.thy @@ -1,2318 +1,2318 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Bit values as reversed lists of bools\ theory Reversed_Bit_Lists imports "HOL-Library.Word" Typedef_Morphisms Least_significant_bit Most_significant_bit Even_More_List "HOL-Library.Sublist" Aligned begin lemma horner_sum_of_bool_2_concat: \horner_sum of_bool 2 (concat (map (\x. map (bit x) [0.. for ws :: \'a::len word list\ proof (induction ws) case Nil then show ?case by simp next case (Cons w ws) moreover have \horner_sum of_bool 2 (map (bit w) [0.. proof transfer fix k :: int have \map (\n. n < LENGTH('a) \ bit k n) [0.. by simp then show \horner_sum of_bool 2 (map (\n. n < LENGTH('a) \ bit k n) [0.. by (simp only: horner_sum_bit_eq_take_bit) qed ultimately show ?case by (simp add: horner_sum_append) qed subsection \Implicit augmentation of list prefixes\ primrec takefill :: "'a \ nat \ 'a list \ 'a list" where Z: "takefill fill 0 xs = []" | Suc: "takefill fill (Suc n) xs = (case xs of [] \ fill # takefill fill n xs | y # ys \ y # takefill fill n ys)" lemma nth_takefill: "m < n \ takefill fill n l ! m = (if m < length l then l ! m else fill)" apply (induct n arbitrary: m l) apply clarsimp apply clarsimp apply (case_tac m) apply (simp split: list.split) apply (simp split: list.split) done lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill" by (induct n arbitrary: l) (auto split: list.split) lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill" by (simp add: takefill_alt replicate_add [symmetric]) lemma takefill_le': "n = m + k \ takefill x m (takefill x n l) = takefill x m l" by (induct m arbitrary: l n) (auto split: list.split) lemma length_takefill [simp]: "length (takefill fill n l) = n" by (simp add: takefill_alt) lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" by (induct k arbitrary: w n) (auto split: list.split) lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" by (induct k arbitrary: w) (auto split: list.split) lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" by (auto simp: le_iff_add takefill_le') lemma take_takefill [simp]: "m \ n \ take m (takefill fill n w) = takefill fill m w" by (auto simp: le_iff_add take_takefill') lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" by (induct xs) auto lemma takefill_same': "l = length xs \ takefill fill l xs = xs" by (induct xs arbitrary: l) auto lemmas takefill_same [simp] = takefill_same' [OF refl] lemma tf_rev: "n + k = m + length bl \ takefill x m (rev (takefill y n bl)) = rev (takefill y m (rev (takefill x k (rev bl))))" apply (rule nth_equalityI) apply (auto simp add: nth_takefill rev_nth) apply (rule_tac f = "\n. bl ! n" in arg_cong) apply arith done lemma takefill_minus: "0 < n \ takefill fill (Suc (n - 1)) w = takefill fill n w" by auto lemmas takefill_Suc_cases = list.cases [THEN takefill.Suc [THEN trans]] lemmas takefill_Suc_Nil = takefill_Suc_cases (1) lemmas takefill_Suc_Cons = takefill_Suc_cases (2) lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] takefill_minus [symmetric, THEN trans]] lemma takefill_numeral_Nil [simp]: "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" by (simp add: numeral_eq_Suc) lemma takefill_numeral_Cons [simp]: "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" by (simp add: numeral_eq_Suc) subsection \Range projection\ definition bl_of_nth :: "nat \ (nat \ 'a) \ 'a list" where "bl_of_nth n f = map f (rev [0.. rev (bl_of_nth n f) ! m = f m" by (simp add: bl_of_nth_def rev_map) lemma bl_of_nth_inj: "(\k. k < n \ f k = g k) \ bl_of_nth n f = bl_of_nth n g" by (simp add: bl_of_nth_def) lemma bl_of_nth_nth_le: "n \ length xs \ bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" apply (induct n arbitrary: xs) apply clarsimp apply clarsimp apply (rule trans [OF _ hd_Cons_tl]) apply (frule Suc_le_lessD) apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric]) apply (subst hd_drop_conv_nth) apply force apply simp_all apply (rule_tac f = "\n. drop n xs" in arg_cong) apply simp done lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs" by (simp add: bl_of_nth_nth_le) subsection \More\ definition rotater1 :: "'a list \ 'a list" where "rotater1 ys = (case ys of [] \ [] | x # xs \ last ys # butlast ys)" definition rotater :: "nat \ 'a list \ 'a list" where "rotater n = rotater1 ^^ n" lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" by (cases l) (auto simp: rotater1_def) lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" apply (unfold rotater1_def) apply (cases "l") apply (case_tac [2] "list") apply auto done lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" by (cases l) (auto simp: rotater1_def) lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" by (induct n) (auto simp: rotater_def intro: rotater1_rev') lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" using rotater_rev' [where xs = "rev ys"] by simp lemma rotater_drop_take: "rotater n xs = drop (length xs - n mod length xs) xs @ take (length xs - n mod length xs) xs" by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" unfolding rotater_def by auto lemma nth_rotater: \rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\ if \n < length xs\ using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) lemma nth_rotater1: \rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\ if \n < length xs\ using that nth_rotater [of n xs 1] by simp lemma rotate_inv_plus [rule_format]: "\k. k = m + n \ rotater k (rotate n xs) = rotater m xs \ rotate k (rotater n xs) = rotate m xs \ rotater n (rotate k xs) = rotate m xs \ rotate n (rotater k xs) = rotater m xs" by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] lemma rotate_gal: "rotater n xs = ys \ rotate n ys = xs" by auto lemma rotate_gal': "ys = rotater n xs \ xs = rotate n ys" by auto lemma length_rotater [simp]: "length (rotater n xs) = length xs" by (simp add : rotater_rev) lemma rotate_eq_mod: "m mod length xs = n mod length xs \ rotate m xs = rotate n xs" apply (rule box_equals) defer apply (rule rotate_conv_mod [symmetric])+ apply simp done lemma restrict_to_left: "x = y \ x = z \ y = z" by simp lemmas rotate_eqs = trans [OF rotate0 [THEN fun_cong] id_apply] rotate_rotate [symmetric] rotate_id rotate_conv_mod rotate_eq_mod lemmas rrs0 = rotate_eqs [THEN restrict_to_left, simplified rotate_gal [symmetric] rotate_gal' [symmetric]] lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] lemmas rotater_eqs = rrs1 [simplified length_rotater] lemmas rotater_0 = rotater_eqs (1) lemmas rotater_add = rotater_eqs (2) lemma butlast_map: "xs \ [] \ butlast (map f xs) = map f (butlast xs)" by (induct xs) auto lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" by (cases xs) (auto simp: rotater1_def last_map butlast_map) lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" by (induct n) (auto simp: rotater_def rotater1_map) lemma but_last_zip [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (zip xs ys) = (last xs, last ys) \ butlast (zip xs ys) = zip (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma but_last_map2 [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (map2 f xs ys) = f (last xs) (last ys) \ butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma rotater1_zip: "length xs = length ys \ rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" apply (unfold rotater1_def) apply (cases xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ done lemma rotater1_map2: "length xs = length ys \ rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" by (simp add: rotater1_map rotater1_zip) lemmas lrth = box_equals [OF asm_rl length_rotater [symmetric] length_rotater [symmetric], THEN rotater1_map2] lemma rotater_map2: "length xs = length ys \ rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" by (induct n) (auto intro!: lrth) lemma rotate1_map2: "length xs = length ys \ rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" by (cases xs; cases ys) auto lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] length_rotate [symmetric], THEN rotate1_map2] lemma rotate_map2: "length xs = length ys \ rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" by (induct n) (auto intro!: lth) subsection \Explicit bit representation of \<^typ>\int\\ primrec bl_to_bin_aux :: "bool list \ int \ int" where Nil: "bl_to_bin_aux [] w = w" | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" definition bl_to_bin :: "bool list \ int" where "bl_to_bin bs = bl_to_bin_aux bs 0" primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" where Z: "bin_to_bl_aux 0 w bl = bl" | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" definition bin_to_bl :: "nat \ int \ bool list" where "bin_to_bl n w = bin_to_bl_aux n w []" lemma bin_to_bl_aux_zero_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" by (cases n) auto lemma bin_to_bl_aux_minus1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" by (cases n) auto lemma bin_to_bl_aux_one_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" by (cases n) auto lemma bin_to_bl_aux_Bit0_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" by (cases n) simp_all lemma bin_to_bl_aux_Bit1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" by (cases n) simp_all lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" by (induct bs arbitrary: w) auto lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" by (induct n arbitrary: w bs) auto lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" by (simp add: bin_to_bl_def bin_to_bl_aux_append) lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" by (auto simp: bin_to_bl_def) lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (induct n arbitrary: w bs) auto lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" by (simp add: bin_to_bl_def size_bin_to_bl_aux) lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" apply (induct bs arbitrary: w n) apply auto apply (simp_all only: add_Suc [symmetric]) apply (auto simp add: bin_to_bl_def) done lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" unfolding bl_to_bin_def apply (rule box_equals) apply (rule bl_bin_bl') prefer 2 apply (rule bin_to_bl_aux.Z) apply simp done lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" apply (rule_tac box_equals) defer apply (rule bl_bin_bl) apply (rule bl_bin_bl) apply simp done lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" by (auto simp: bl_to_bin_def) lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" by (auto simp: bl_to_bin_def) lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" by (simp add: bin_to_bl_def bin_to_bl_zero_aux) lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" by (auto simp: bin_to_bl_def bin_bl_bin') lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (bintrunc m w) = bin_to_bl n w" by (auto intro: bl_to_bin_inj) lemma bin_to_bl_aux_bintr: "bin_to_bl_aux n (bintrunc m bin) bl = replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" apply (induct n arbitrary: m bin bl) apply clarsimp apply clarsimp apply (case_tac "m") apply (clarsimp simp: bin_to_bl_zero_aux) apply (erule thin_rl) apply (induct_tac n) apply (auto simp add: take_bit_Suc) done lemma bin_to_bl_bintr: "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" by (induct n) auto lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (fact size_bin_to_bl_aux) lemma len_bin_to_bl: "length (bin_to_bl n w) = n" by (fact size_bin_to_bl) (* FIXME: duplicate *) lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" by (induction bs arbitrary: w) (simp_all add: bin_sign_def) lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" by (simp add: bl_to_bin_def sign_bl_bin') lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) lemma bin_nth_of_bl_aux: "bin_nth (bl_to_bin_aux bl w) n = (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" apply (induction bl arbitrary: w) apply simp_all apply safe apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) done lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" by (simp add: bl_to_bin_def bin_nth_of_bl_aux) lemma bin_nth_bl: "n < m \ bin_nth w n = nth (rev (bin_to_bl m w)) n" apply (induct n arbitrary: m w) apply clarsimp apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt bit_Suc) done lemma nth_bin_to_bl_aux: "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = (if n < m then bit w (m - 1 - n) else bl ! (n - m))" apply (induction bl arbitrary: w) apply simp_all apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) apply (metis One_nat_def Suc_pred add_diff_cancel_left' add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) done lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" by (simp add: bin_to_bl_def nth_bin_to_bl_aux) lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" apply (rule nth_equalityI) apply simp apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) done lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" by (simp add: takefill_bintrunc) lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (subst mult_2 [of \2 ^ length bs\]) apply (simp only: add.assoc) apply (rule pos_add_strict) apply simp_all done qed lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" proof (induct bs) case Nil then show ?case by simp next case (Cons b bs) with bl_to_bin_lt2p_aux[where w=1] show ?case by (simp add: bl_to_bin_def) qed lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" by (metis bin_bl_bin bintr_lt2p bl_bin_bl) lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (rule add_le_imp_le_left [of \2 ^ length bs\]) apply (rule add_increasing) apply simp_all done qed lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" apply (unfold bl_to_bin_def) apply (rule xtrans(4)) apply (rule bl_to_bin_ge2p_aux) apply simp done lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" apply (unfold bin_to_bl_def) apply (cases n, clarsimp) apply clarsimp apply (auto simp add: bin_to_bl_aux_alt) done lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp lemma butlast_rest_bl2bin_aux: "bl \ [] \ bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" by (induct bl arbitrary: w) auto lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) lemma trunc_bl2bin_aux: "bintrunc m (bl_to_bin_aux bl w) = bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" proof (induct bl arbitrary: w) case Nil show ?case by simp next case (Cons b bl) show ?case proof (cases "m - length bl") case 0 then have "Suc (length bl) - m = Suc (length bl - m)" by simp with Cons show ?thesis by simp next case (Suc n) then have "m - Suc (length bl) = n" by simp with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) qed qed lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" by (simp add: bl_to_bin_def trunc_bl2bin_aux) lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" by (simp add: trunc_bl2bin) lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" apply (rule trans) prefer 2 apply (rule trunc_bl2bin [symmetric]) apply (cases "k \ length bl") apply auto done lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) done lemma last_bin_last': "size xs > 0 \ last xs \ bin_last (bl_to_bin_aux xs w)" by (induct xs arbitrary: w) auto lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (bl_to_bin xs)" unfolding bl_to_bin_def by (erule last_bin_last') lemma bin_last_last: "bin_last w \ last (bin_to_bl (Suc n) w)" by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) lemma drop_bin2bl_aux: "drop m (bin_to_bl_aux n bin bs) = bin_to_bl_aux (n - m) bin (drop (m - n) bs)" apply (induction n arbitrary: m bin bs) apply auto apply (case_tac "m \ n") apply (auto simp add: not_le Suc_diff_le) apply (case_tac "m - n") apply auto apply (use Suc_diff_Suc in fastforce) done lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" by (simp add: bin_to_bl_def drop_bin2bl_aux) lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" apply (induct m arbitrary: w bs) apply clarsimp apply clarsimp apply (simp add: bin_to_bl_aux_alt) apply (simp add: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) done lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" apply (induct n arbitrary: b c) apply clarsimp apply (clarsimp simp: Let_def split: prod.split_asm) apply (simp add: bin_to_bl_def) apply (simp add: take_bin2bl_lem drop_bit_Suc) done lemma bin_to_bl_drop_bit: "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" using bin_split_take by simp lemma bin_split_take1: "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" using bin_split_take by simp lemma bl_bin_bl_rep_drop: "bin_to_bl n (bl_to_bin bl) = replicate (n - length bl) False @ drop (length bl - n) bl" by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) lemma bl_to_bin_aux_cat: "bl_to_bin_aux bs (bin_cat w nv v) = bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" by (rule bit_eqI) (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) lemma bin_to_bl_aux_cat: "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" using bl_to_bin_aux_cat [where nv = "0" and v = "0"] by (simp add: bl_to_bin_def [symmetric]) lemma bin_to_bl_cat: "bin_to_bl (nv + nw) (bin_cat v nw w) = bin_to_bl_aux nv v (bin_to_bl nw w)" by (simp add: bin_to_bl_def bin_to_bl_aux_cat) lemmas bl_to_bin_aux_app_cat = trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] lemmas bin_to_bl_aux_cat_app = trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] lemma bl_to_bin_app_cat: "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) lemma bin_to_bl_cat_app: "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" by (simp add: bl_to_bin_app_cat) lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" apply (unfold bl_to_bin_def) apply (induct n) apply simp apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) apply simp done lemma bin_exhaust: "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int apply (cases \even bin\) apply (auto elim!: evenE oddE) apply fastforce apply fastforce done primrec rbl_succ :: "bool list \ bool list" where Nil: "rbl_succ Nil = Nil" | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" primrec rbl_pred :: "bool list \ bool list" where Nil: "rbl_pred Nil = Nil" | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" primrec rbl_add :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_add Nil x = Nil" | Cons: "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" primrec rbl_mult :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_mult Nil x = Nil" | Cons: "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in if y then rbl_add ws x else ws)" lemma size_rbl_pred: "length (rbl_pred bl) = length bl" by (induct bl) auto lemma size_rbl_succ: "length (rbl_succ bl) = length bl" by (induct bl) auto lemma size_rbl_add: "length (rbl_add bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) lemmas rbl_sizes [simp] = size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult lemmas rbl_Nils = rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_add_take2: "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def rbl_add_app2) done lemma rbl_mult_take2: "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" apply (rule trans) apply (rule rbl_mult_app2 [symmetric]) apply simp apply (rule_tac f = "rbl_mult bla" in arg_cong) apply (rule append_take_drop_id) done lemma rbl_add_split: "P (rbl_add (y # ys) (x # xs)) = (\ws. length ws = length ys \ ws = rbl_add ys xs \ (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ (\ y \ P (x # ws)))" by (cases y) (auto simp: Let_def) lemma rbl_mult_split: "P (rbl_mult (y # ys) xs) = (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" by (auto simp: Let_def) lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" proof (unfold bin_to_bl_def, induction n arbitrary: bin) case 0 then show ?case by simp next case (Suc n) obtain b k where \bin = of_bool b + 2 * k\ using bin_exhaust by blast moreover have \(2 * k - 1) div 2 = k - 1\ using even_succ_div_2 [of \2 * (k - 1)\] by simp ultimately show ?case using Suc [of \bin div 2\] by simp (auto simp add: bin_to_bl_aux_alt) qed lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" apply (unfold bin_to_bl_def) apply (induction n arbitrary: bin) apply simp_all apply (case_tac bin rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt ac_simps) done lemma rbl_add: "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina + binb))" apply (unfold bin_to_bl_def) apply (induct n) apply simp apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (case_tac b) apply (case_tac [!] "ba") apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) done lemma rbl_add_long: "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rev (bin_to_bl n (bina + binb))" apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) apply simp done lemma rbl_mult_gt1: "m \ length bl \ rbl_mult bl (rev (bin_to_bl m binb)) = rbl_mult bl (rev (bin_to_bl (length bl) binb))" apply (rule trans) apply (rule rbl_mult_take2 [symmetric]) apply simp_all apply (rule_tac f = "rbl_mult bl" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) done lemma rbl_mult_gt: "m > n \ rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" by (auto intro: trans [OF rbl_mult_gt1]) lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) lemma rbl_mult: "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina * binb))" apply (induct n arbitrary: bina binb) apply simp_all apply (unfold bin_to_bl_def) apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt) apply (simp_all add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) done lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" by (simp add: length_concat comp_def sum_list_triv) lemma bin_cat_foldl_lem: "foldl (\u. bin_cat u n) x xs = bin_cat x (size xs * n) (foldl (\u. bin_cat u n) y xs)" apply (induct xs arbitrary: x) apply simp apply (simp (no_asm)) apply (frule asm_rl) apply (drule meta_spec) apply (erule trans) apply (drule_tac x = "bin_cat y n a" in meta_spec) apply (simp add: bin_cat_assoc_sym) done lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" apply (unfold bin_rcat_eq_foldl) apply (rule sym) apply (induct wl) apply (auto simp add: bl_to_bin_append) apply (simp add: bl_to_bin_aux_alt sclem) apply (simp add: bin_cat_foldl_lem [symmetric]) done lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \ bs \ [] \ last bs" by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)" by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) lemma bl_xor_aux_bin: "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" apply (induction n arbitrary: v w bs cs) apply auto apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp done lemma bl_or_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" by (induct n arbitrary: v w bs cs) simp_all lemma bl_and_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" by (induction n arbitrary: v w bs cs) simp_all lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" by (induct n arbitrary: w cs) auto lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" by (simp add: bin_to_bl_def bl_not_aux_bin) lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" by (simp add: bin_to_bl_def bl_and_aux_bin) lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" by (simp add: bin_to_bl_def bl_or_aux_bin) lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" using bl_xor_aux_bin by (simp add: bin_to_bl_def) subsection \Type \<^typ>\'a word\\ lift_definition of_bl :: \bool list \ 'a::len word\ is bl_to_bin . lift_definition to_bl :: \'a::len word \ bool list\ is \bin_to_bl LENGTH('a)\ by (simp add: bl_to_bin_inj) lemma to_bl_eq: \to_bl w = bin_to_bl (LENGTH('a)) (uint w)\ for w :: \'a::len word\ by transfer simp lemma bit_of_bl_iff [bit_simps]: \bit (of_bl bs :: 'a word) n \ rev bs ! n \ n < LENGTH('a::len) \ n < length bs\ by transfer (simp add: bin_nth_of_bl ac_simps) lemma rev_to_bl_eq: \rev (to_bl w) = map (bit w) [0.. for w :: \'a::len word\ apply (rule nth_equalityI) apply (simp add: to_bl.rep_eq) apply (simp add: bin_nth_bl bit_word.rep_eq to_bl.rep_eq) done lemma to_bl_eq_rev: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ using rev_to_bl_eq [of w] apply (subst rev_is_rev_conv [symmetric]) apply (simp add: rev_map) done lemma of_bl_rev_eq: \of_bl (rev bs) = horner_sum of_bool 2 bs\ apply (rule bit_word_eqI) apply (simp add: bit_of_bl_iff) apply transfer apply (simp add: bit_horner_sum_bit_iff ac_simps) done lemma of_bl_eq: \of_bl bs = horner_sum of_bool 2 (rev bs)\ using of_bl_rev_eq [of \rev bs\] by simp lemma bshiftr1_eq: \bshiftr1 b w = of_bl (b # butlast (to_bl w))\ apply transfer apply auto apply (subst bl_to_bin_app_cat [of \[True]\, simplified]) apply simp apply (metis One_nat_def add.commute bin_bl_bin bin_last_bl_to_bin bin_rest_bl_to_bin butlast_bin_rest concat_bit_eq last.simps list.distinct(1) list.size(3) list.size(4) odd_iff_mod_2_eq_one plus_1_eq_Suc power_Suc0_right push_bit_of_1 size_bin_to_bl take_bit_eq_mod trunc_bl2bin_len) apply (simp add: butlast_rest_bl2bin) done lemma length_to_bl_eq: \length (to_bl w) = LENGTH('a)\ for w :: \'a::len word\ by transfer simp lemma word_rotr_eq: \word_rotr n w = of_bl (rotater n (to_bl w))\ apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff rotater_rev length_to_bl_eq nth_rotate rev_to_bl_eq ac_simps) done done lemma word_rotl_eq: \word_rotl n w = of_bl (rotate n (to_bl w))\ proof - have \rotate n (to_bl w) = rev (rotater n (rev (to_bl w)))\ by (simp add: rotater_rev') then show ?thesis apply (simp add: word_rotl_eq_word_rotr bit_of_bl_iff length_to_bl_eq rev_to_bl_eq) apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff nth_rotater) done done qed lemma to_bl_def': "(to_bl :: 'a::len word \ bool list) = bin_to_bl (LENGTH('a)) \ uint" by transfer (simp add: fun_eq_iff) \ \type definitions theorem for in terms of equivalent bool list\ lemma td_bl: "type_definition (to_bl :: 'a::len word \ bool list) of_bl {bl. length bl = LENGTH('a)}" apply (standard; transfer) apply (auto dest: sym) done interpretation word_bl: type_definition "to_bl :: 'a::len word \ bool list" of_bl "{bl. length bl = LENGTH('a::len)}" by (fact td_bl) lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] lemma word_size_bl: "size w = size (to_bl w)" by (auto simp: word_size) lemma to_bl_use_of_bl: "to_bl w = bl \ w = of_bl bl \ length bl = length (to_bl w)" by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" for x :: "'a::len word" unfolding word_bl_Rep' by (rule len_gt_0) lemma bl_not_Nil [iff]: "to_bl x \ []" for x :: "'a::len word" by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" for x :: "'a::len word" by (fact length_bl_gt_0 [THEN gr_implies_not0]) lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" apply transfer apply (auto simp add: bin_sign_def) using bin_sign_lem bl_sbin_sign apply fastforce using bin_sign_lem bl_sbin_sign apply force done lemma of_bl_drop': "lend = length bl - LENGTH('a::len) \ of_bl (drop lend bl) = (of_bl bl :: 'a word)" by transfer (simp flip: trunc_bl2bin) lemma test_bit_of_bl: "(of_bl bl::'a::len word) !! n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" by transfer (simp add: bin_nth_of_bl ac_simps) lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" by transfer simp lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" by transfer simp lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin" by (auto simp: uint_bl word_ubin.eq_norm word_size) lemma to_bl_numeral [simp]: "to_bl (numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (numeral bin)" unfolding word_numeral_alt by (rule to_bl_of_bin) lemma to_bl_neg_numeral [simp]: "to_bl (- numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (- numeral bin)" unfolding word_neg_numeral_alt by (rule to_bl_of_bin) lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x" for x :: "'a::len word" by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) lemma ucast_bl: "ucast w = of_bl (to_bl w)" by transfer simp lemma ucast_down_bl: \(ucast :: 'a::len word \ 'b::len word) (of_bl bl) = of_bl bl\ if \is_down (ucast :: 'a::len word \ 'b::len word)\ using that by transfer simp lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" by transfer (simp add: bl_to_bin_app_cat) lemma ucast_of_bl_up: \ucast (of_bl bl :: 'a::len word) = of_bl bl\ if \size bl \ size (of_bl bl :: 'a::len word)\ using that apply transfer apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff) apply (subst (asm) trunc_bl2bin_len [symmetric]) apply (auto simp only: bit_take_bit_iff) done lemma word_rev_tf: "to_bl (of_bl bl::'a::len word) = rev (takefill False (LENGTH('a)) (rev bl))" by transfer (simp add: bl_bin_bl_rtf) lemma word_rep_drop: "to_bl (of_bl bl::'a::len word) = replicate (LENGTH('a) - length bl) False @ drop (length bl - LENGTH('a)) bl" by (simp add: word_rev_tf takefill_alt rev_take) lemma to_bl_ucast: "to_bl (ucast (w::'b::len word) ::'a::len word) = replicate (LENGTH('a) - LENGTH('b)) False @ drop (LENGTH('b) - LENGTH('a)) (to_bl w)" apply (unfold ucast_bl) apply (rule trans) apply (rule word_rep_drop) apply simp done lemma ucast_up_app: \to_bl (ucast w :: 'b::len word) = replicate n False @ (to_bl w)\ if \source_size (ucast :: 'a word \ 'b word) + n = target_size (ucast :: 'a word \ 'b word)\ for w :: \'a::len word\ using that by (auto simp add : source_size target_size to_bl_ucast) lemma ucast_down_drop [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ to_bl (uc w) = drop n (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma scast_down_drop [OF refl]: "sc = scast \ source_size sc = target_size sc + n \ to_bl (sc w) = drop n (to_bl w)" apply (subgoal_tac "sc = ucast") apply safe apply simp apply (erule ucast_down_drop) apply (rule down_cast_same [symmetric]) apply (simp add : source_size target_size is_down) done lemma word_0_bl [simp]: "of_bl [] = 0" by transfer simp lemma word_1_bl: "of_bl [True] = 1" by transfer (simp add: bl_to_bin_def) lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" by transfer (simp add: bl_to_bin_rep_False) lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False" by (simp add: uint_bl word_size bin_to_bl_zero) \ \links with \rbl\ operations\ lemma word_succ_rbl: "to_bl w = bl \ to_bl (word_succ w) = rev (rbl_succ (rev bl))" by transfer (simp add: rbl_succ) lemma word_pred_rbl: "to_bl w = bl \ to_bl (word_pred w) = rev (rbl_pred (rev bl))" by transfer (simp add: rbl_pred) lemma word_add_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_add) done lemma word_mult_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_mult) done lemma rtb_rbl_ariths: "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v * w)) = rbl_mult ys xs" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v + w)) = rbl_add ys xs" by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) lemma of_bl_length_less: \(of_bl x :: 'a::len word) < 2 ^ k\ if \length x = k\ \k < LENGTH('a)\ proof - from that have \length x < LENGTH('a)\ by simp then have \(of_bl x :: 'a::len word) < 2 ^ length x\ apply (simp add: of_bl_eq) apply transfer apply (simp add: take_bit_horner_sum_bit_eq) apply (subst length_rev [symmetric]) apply (simp only: horner_sum_of_bool_2_less) done with that show ?thesis by simp qed lemma word_eq_rbl_eq: "x = y \ rev (to_bl x) = rev (to_bl y)" by simp lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" by transfer (simp add: bl_not_bin) lemma bl_word_xor: "to_bl (v XOR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_xor_bin) lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_or_bin) lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_and_bin) lemma bin_nth_uint': "bin_nth (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) apply (fast dest!: bin_nth_bl)+ done lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] lemma test_bit_bl: "w !! n \ rev (to_bl w) ! n \ n < size w" by transfer (auto simp add: bin_nth_bl) lemma to_bl_nth: "n < size w \ to_bl w ! n = w !! (size w - Suc n)" by (simp add: word_size rev_nth test_bit_bl) lemma map_bit_interval_eq: \map (bit w) [0.. for w :: \'a::len word\ proof (rule nth_equalityI) show \length (map (bit w) [0.. by simp fix m assume \m < length (map (bit w) [0.. then have \m < n\ by simp then have \bit w m \ takefill False n (rev (to_bl w)) ! m\ by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size test_bit_word_eq dest: bit_imp_le_length) with \m < n \show \map (bit w) [0.. takefill False n (rev (to_bl w)) ! m\ by simp qed lemma to_bl_unfold: \to_bl w = rev (map (bit w) [0.. for w :: \'a::len word\ by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def) lemma nth_rev_to_bl: \rev (to_bl w) ! n \ bit w n\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold) lemma nth_to_bl: \to_bl w ! n \ bit w (LENGTH('a) - Suc n)\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold rev_nth) lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" by (auto simp: of_bl_def bl_to_bin_rep_F) lemma [code abstract]: \Word.the_int (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\ apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq) apply transfer apply simp done lemma [code]: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ by (fact to_bl_eq_rev) lemma word_reverse_eq_of_bl_rev_to_bl: \word_reverse w = of_bl (rev (to_bl w))\ by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) lemmas word_reverse_no_def [simp] = word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_or rev_map) lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_and rev_map) lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_xor rev_map) lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" by (simp add: bl_word_not rev_map) lemma bshiftr1_numeral [simp]: \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ by (simp add: bshiftr1_eq) lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" by transfer (simp add: bl_to_bin_append) lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" for w :: "'a::len word" proof - have "shiftl1 w = shiftl1 (of_bl (to_bl w))" by simp also have "\ = of_bl (to_bl w @ [False])" by (rule shiftl1_of_bl) finally show ?thesis . qed lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" for w :: "'a::len word" by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) \ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) lemma shiftr1_bl: \shiftr1 w = of_bl (butlast (to_bl w))\ proof (rule bit_word_eqI) fix n assume \n < LENGTH('a)\ show \bit (shiftr1 w) n \ bit (of_bl (butlast (to_bl w)) :: 'a word) n\ proof (cases \n = LENGTH('a) - 1\) case True then show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff) next case False with \n < LENGTH('a)\ have \n < LENGTH('a) - 1\ by simp with \n < LENGTH('a)\ show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast word_size test_bit_word_eq to_bl_nth) qed qed lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" for w :: "'a::len word" by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) \ \Generalized version of \bl_shiftr1\. Maybe this one should replace it?\ lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" apply (rule word_bl.Abs_inverse') apply (simp del: butlast.simps) apply (simp add: shiftr1_bl of_bl_def) done lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" for w :: "'a::len word" proof (rule nth_equalityI) fix n assume \n < length (to_bl (sshiftr1 w))\ then have \n < LENGTH('a)\ by simp then show \to_bl (sshiftr1 w) ! n \ (hd (to_bl w) # butlast (to_bl w)) ! n\ apply (cases n) apply (simp_all add: to_bl_nth word_size hd_conv_nth test_bit_eq_bit bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) done qed simp lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) apply (rule butlast_take [THEN trans]) apply (auto simp: word_size) done lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (simp_all add: word_size sshiftr_eq) apply (rule nth_equalityI) apply (simp_all add: word_size nth_to_bl bit_signed_drop_bit_iff) done lemma take_shiftr: "n \ size w \ take n (to_bl (w >> n)) = replicate n False" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size) apply (rule take_butlast [THEN trans]) apply (auto simp: word_size) done lemma take_sshiftr': "n \ size w \ hd (to_bl (w >>> n)) = hd (to_bl w) \ take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" for w :: "'a::len word" apply (auto simp add: sshiftr_eq hd_bl_sign_sint bin_sign_def not_le word_size sint_signed_drop_bit_eq) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) done lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] lemma atd_lem: "take n xs = t \ drop n xs = d \ xs = t @ d" by (auto intro: append_take_drop_id [symmetric]) lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same) lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" for w :: "'a::len word" proof - have "w << n = of_bl (to_bl w) << n" by simp also have "\ = of_bl (to_bl w @ replicate n False)" by (rule shiftl_of_bl) finally show ?thesis . qed lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" by (simp add: shiftl_bl word_rep_drop word_size) lemma shiftr1_bl_of: "length bl \ LENGTH('a) \ shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" by transfer (simp add: butlast_rest_bl2bin trunc_bl2bin) lemma shiftr_bl_of: "length bl \ LENGTH('a) \ (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" apply (unfold shiftr_def) apply (induct n) apply clarsimp apply clarsimp apply (subst shiftr1_bl_of) apply simp apply (simp add: butlast_take) done lemma shiftr_bl: "x >> n \ of_bl (take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp lemma aligned_bl_add_size [OF refl]: "size x - n = m \ n \ size x \ drop m (to_bl x) = replicate n False \ take m (to_bl y) = replicate m False \ to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \'a::len word\ apply (subgoal_tac "x AND y = 0") prefer 2 apply (rule word_bl.Rep_eqD) apply (simp add: bl_word_and) apply (rule align_lem_and [THEN trans]) apply (simp_all add: word_size)[5] apply simp apply (subst word_plus_and_or [symmetric]) apply (simp add : bl_word_or) apply (rule align_lem_or) apply (simp_all add: word_size) done lemma mask_bl: "mask n = of_bl (replicate n True)" by (auto simp add : test_bit_of_bl word_size intro: word_eqI) lemma bl_and_mask': "to_bl (w AND mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: to_bl_nth word_size) apply (auto simp add: word_size test_bit_bl nth_append rev_nth) done lemma slice1_eq_of_bl: \(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\ for w :: \'a::len word\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('b)\ show \bit (slice1 n w :: 'b::len word) m \ bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\ by (cases \m \ n\; cases \LENGTH('a) \ n\) (auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps) qed lemma slice1_no_bin [simp]: "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *) lemma slice_no_bin [simp]: "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n) (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice_def) (* TODO: neg_numeral *) lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt) lemmas slice_take = slice_take' [unfolded word_size] \ \shiftr to a word of the same size is just slice, slice is just shiftr then ucast\ lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] lemma slice1_down_alt': "sl = slice1 n w \ fs = size sl \ fs + k = n \ to_bl sl = takefill False fs (drop k (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop) using drop_takefill apply force done lemma slice1_up_alt': "sl = slice1 n w \ fs = size sl \ fs = n + k \ to_bl sl = takefill False fs (replicate k False @ (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop flip: takefill_append) apply (metis diff_add_inverse) done lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] lemmas slice1_up_alts = le_add_diff_inverse [symmetric, THEN su1] le_add_diff_inverse2 [symmetric, THEN su1] lemma slice1_tf_tf': "to_bl (slice1 n w :: 'a::len word) = rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" unfolding slice1_eq_of_bl by (rule word_rev_tf) lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] lemma revcast_eq_of_bl: \(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\ for w :: \'a::len word\ by (simp add: revcast_def slice1_eq_of_bl) lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w lemma to_bl_revcast: "to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)" apply (rule nth_equalityI) apply simp apply (cases \LENGTH('a) \ LENGTH('b)\) apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff) done lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" apply (rule bit_word_eqI) apply (simp add: bit_word_cat_iff bit_of_bl_iff nth_append not_less nth_rev_to_bl) apply (meson bit_word.rep_eq less_diff_conv2 nth_rev_to_bl) done lemma of_bl_append: "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" apply transfer apply (simp add: bl_to_bin_app_cat bin_cat_num) done lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" by (cases x) simp_all lemma word_split_bl': "std = size c - size b \ (word_split c = (a, b)) \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c)))" apply (simp add: word_split_def) apply transfer apply (cases \LENGTH('b) \ LENGTH('a)\) apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\] min_absorb2) done lemma word_split_bl: "std = size c - size b \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c))) \ word_split c = (a, b)" apply (rule iffI) defer apply (erule (1) word_split_bl') apply (case_tac "word_split c") apply (auto simp add: word_size) apply (frule word_split_bl' [rotated]) apply (auto simp add: word_size) done lemma word_split_bl_eq: "(word_split c :: ('c::len word \ 'd::len word)) = (of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)), of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" for c :: "'a::len word" apply (rule word_split_bl [THEN iffD1]) apply (unfold word_size) apply (rule refl conjI)+ done lemma word_rcat_bl: \word_rcat wl = of_bl (concat (map to_bl wl))\ proof - define ws where \ws = rev wl\ moreover have \word_rcat (rev ws) = of_bl (concat (map to_bl (rev ws)))\ apply (simp add: word_rcat_def of_bl_eq rev_concat rev_map comp_def rev_to_bl_eq flip: horner_sum_of_bool_2_concat) apply transfer apply simp done ultimately show ?thesis by simp qed lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" by (induct wl) (auto simp: word_size) lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] lemma nth_rcat_lem: "n < length (wl::'a word list) * LENGTH('a::len) \ rev (concat (map to_bl wl)) ! n = rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" apply (induct wl) apply clarsimp apply (clarsimp simp add : nth_append size_rcat_lem) apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less) apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0) done lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" for x :: "'a::comm_monoid_add" by (induct xs arbitrary: x) (auto simp: add.assoc) lemmas word_cat_bl_no_bin [simp] = word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] for a b (* FIXME: negative numerals, 0 and 1 *) lemmas word_split_bl_no_bin [simp] = word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c lemmas word_rot_defs = word_roti_eq_word_rotr_word_rotl word_rotr_eq word_rotl_eq lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" by (simp add: word_rotl_eq to_bl_use_of_bl) lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] lemmas word_rotl_eqs = blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" by (simp add: word_rotr_eq to_bl_use_of_bl) lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] lemmas word_rotr_eqs = brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] declare word_rotr_eqs (1) [simp] declare word_rotl_eqs (1) [simp] lemmas abl_cong = arg_cong [where f = "of_bl"] locale word_rotate begin lemmas word_rot_defs' = to_bl_rotl to_bl_rotr lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map end lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, simplified word_bl_Rep'] lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, simplified word_bl_Rep'] lemma bl_word_roti_dt': "n = nat ((- i) mod int (size (w :: 'a::len word))) \ to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" apply (unfold word_roti_eq_word_rotr_word_rotl) apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) apply safe apply (simp add: zmod_zminus1_eq_if) apply safe apply (simp add: nat_mult_distrib) apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj [THEN conjunct2, THEN order_less_imp_le]] nat_mod_distrib) apply (simp add: nat_mod_distrib) done lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt_no_bin' [simp] = word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemmas word_rotl_dt_no_bin' [simp] = word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) -lemma max_word_bl: "to_bl (max_word::'a::len word) = replicate LENGTH('a) True" +lemma max_word_bl: "to_bl (- 1::'a::len word) = replicate LENGTH('a) True" by (fact to_bl_n1) lemma to_bl_mask: "to_bl (mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ replicate (min (LENGTH('a)) n) True" by (simp add: mask_bl word_rep_drop min_def) lemma map_replicate_True: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n True)) = xs" by (induct xs arbitrary: n) auto lemma map_replicate_False: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto lemma bl_and_mask: fixes w :: "'a::len word" and n :: nat defines "n' \ LENGTH('a) - n" shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" proof - note [simp] = map_replicate_True map_replicate_False have "to_bl (w AND mask n) = map2 (\) (to_bl w) (to_bl (mask n::'a::len word))" by (simp add: bl_word_and) also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp also have "map2 (\) \ (to_bl (mask n::'a::len word)) = replicate n' False @ drop n' (to_bl w)" unfolding to_bl_mask n'_def by (subst zip_append) auto finally show ?thesis . qed lemma drop_rev_takefill: "length xs \ n \ drop (n - length xs) (rev (takefill False n (rev xs))) = xs" by (simp add: takefill_alt rev_take) declare bin_to_bl_def [simp] lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append lemma uint_of_bl_is_bl_to_bin_drop: "length (dropWhile Not l) \ LENGTH('a) \ uint (of_bl l :: 'a::len word) = bl_to_bin l" apply transfer apply (simp add: take_bit_eq_mod) apply (rule Divides.mod_less) apply (rule bl_to_bin_ge0) using bl_to_bin_lt2p_drop apply (rule order.strict_trans2) apply simp done corollary uint_of_bl_is_bl_to_bin: "length l\LENGTH('a) \ uint ((of_bl::bool list\ ('a :: len) word) l) = bl_to_bin l" apply(rule uint_of_bl_is_bl_to_bin_drop) using le_trans length_dropWhile_le by blast lemma bin_to_bl_or: "bin_to_bl n (a OR b) = map2 (\) (bin_to_bl n a) (bin_to_bl n b)" using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"] by simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq and_one_eq) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq one_and_eq) lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs AND mask (length xs - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe; simp add: word_size to_bl_nth) done lemma to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" by (rule nth_equalityI) (auto simp add: to_bl_unfold nth_append rev_nth bit_1_iff not_less not_le) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" apply (auto simp add: to_bl_unfold) apply (rule bit_word_eqI) apply auto done lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x AND mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma word_msb_alt: "msb w \ hd (to_bl w)" for w :: "'a::len word" apply (simp add: msb_word_eq) apply (subst hd_conv_nth) apply simp apply (subst nth_to_bl) apply simp apply simp done lemma word_lsb_last: \lsb w \ last (to_bl w)\ for w :: \'a::len word\ using nth_to_bl [of \LENGTH('a) - Suc 0\ w] by (simp add: lsb_odd last_conv_nth) lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" apply (simp add: is_aligned_mask eq_zero_set_bl) apply (clarsimp simp: in_set_conv_nth word_size) apply (simp add: to_bl_nth word_size cong: conj_cong) apply (simp add: diff_diff_less) apply safe apply (case_tac "n \ LENGTH('a)") prefer 2 apply (rule_tac x=i in exI) apply clarsimp apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") apply (erule exE) apply (rule_tac x=j in exI) apply clarsimp apply (thin_tac "w !! t" for t) apply (rule_tac x="i + n - LENGTH('a)" in exI) apply clarsimp apply arith apply (rule_tac x="LENGTH('a) - n + i" in exI) apply clarsimp apply arith done lemma is_aligned_replicate: fixes w::"'a::len word" assumes aligned: "is_aligned w n" and nv: "n \ LENGTH('a)" shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" proof - from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ to_bl (2 ^ n * (of_nat q :: 'a word)) = drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) show ?thesis using aligned by (auto simp: rl elim: is_alignedE) qed lemma is_aligned_drop: fixes w::"'a::len word" assumes "is_aligned w n" "n \ LENGTH('a)" shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" proof - have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" by (rule is_aligned_replicate) fact+ then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp also have "\ = replicate n False" by simp finally show ?thesis . qed lemma less_is_drop_replicate: fixes x::"'a::len word" assumes lt: "x < 2 ^ n" shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" by (metis assms bl_and_mask' less_mask_eq) lemma is_aligned_add_conv: fixes off::"'a::len word" assumes aligned: "is_aligned w n" and offv: "off < 2 ^ n" shows "to_bl (w + off) = (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" proof cases assume nv: "n \ LENGTH('a)" show ?thesis proof (subst aligned_bl_add_size, simp_all only: word_size) show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) from offv show "take (LENGTH('a) - n) (to_bl off) = replicate (LENGTH('a) - n) False" by (subst less_is_drop_replicate, assumption) simp qed fact next assume "\ n \ LENGTH('a)" with offv show ?thesis by (simp add: power_overflow) qed lemma is_aligned_replicateI: "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" apply (simp add: is_aligned_to_bl word_size) apply (subgoal_tac "length addr = LENGTH('a) - n") apply (simp add: replicate_not_True) apply (drule arg_cong [where f=length]) apply simp done lemma to_bl_2p: "n < LENGTH('a) \ to_bl ((2::'a::len word) ^ n) = replicate (LENGTH('a) - Suc n) False @ True # replicate n False" apply (subst shiftl_1 [symmetric]) apply (subst bl_shiftl) apply (simp add: to_bl_1 min_def word_size) done lemma xor_2p_to_bl: fixes x::"'a::len word" shows "to_bl (x XOR 2^n) = (if n < LENGTH('a) then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) else to_bl x)" proof - have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" by simp show ?thesis apply simp apply (rule conjI) apply (clarsimp simp: word_size) apply (simp add: bl_word_xor to_bl_2p) apply (subst x) apply (subst zip_append) apply simp apply (simp add: map_zip_replicate_False_xor drop_minus) apply (auto simp add: word_size nth_w2p intro!: word_eqI) done qed lemma is_aligned_replicateD: "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ \ \xs. to_bl w = xs @ replicate n False \ length xs = size w - n" apply (subst is_aligned_replicate, assumption+) apply (rule exI, rule conjI, rule refl) apply (simp add: word_size) done text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) AND NOT(mask n) = a" apply (simp flip: push_bit_eq_mult subtract_mask(1) take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (simp add: take_bit_push_bit) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (simp flip: push_bit_eq_mult take_bit_eq_mask add: shiftr_eq_drop_bit) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less unsigned_or_eq unsigned_drop_bit_eq unsigned_push_bit_eq bin_to_bl_or simp flip: bin_to_bl_def) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (rule nth_equalityI) apply (auto simp add: nth_bin_to_bl bit_simps rev_nth simp flip: bin_to_bl_def) done (* FIXME: move to Word distribution *) lemma bin_nth_minus_Bit0[simp]: "0 < n \ bin_nth (numeral (num.Bit0 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: "0 < n \ bin_nth (numeral (num.Bit1 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma map_bits_rev_to_bl: "map ((!!) x) [0.. of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" proof - define ys where \ys = rev xs\ have \take_bit (length ys) (horner_sum of_bool 2 ys :: 'a word) = horner_sum of_bool 2 ys\ by transfer (simp add: take_bit_horner_sum_bit_eq min_def) then have \(of_bl (rev ys) :: 'a word) \ mask (length ys)\ by (simp only: of_bl_rev_eq less_eq_mask_iff_take_bit_eq_self) with ys_def show ?thesis by simp qed text\Some auxiliaries for sign-shifting by the entire word length or more\ lemma sshiftr_clamp_pos: assumes "LENGTH('a) \ n" "0 \ sint x" shows "(x::'a::len word) >>> n = 0" apply (rule word_sint.Rep_eqD) apply (unfold sshiftr_div_2n Word.sint_0) apply (rule div_pos_pos_trivial) subgoal using assms(2) . apply (rule order.strict_trans[where b="2 ^ (LENGTH('a) - 1)"]) using sint_lt assms(1) by auto lemma sshiftr_clamp_neg: assumes "LENGTH('a) \ n" "sint x < 0" shows "(x::'a::len word) >>> n = -1" proof - have *: "- (2 ^ n) < sint x" apply (rule order.strict_trans2[where b="- (2 ^ (LENGTH('a) - 1))"]) using assms(1) sint_ge by auto show ?thesis apply (rule word_sint.Rep_eqD) apply (unfold sshiftr_div_2n Word.sint_n1) apply (subst div_minus_minus[symmetric]) apply (rule div_pos_neg_trivial) subgoal using assms(2) by linarith using * by simp qed lemma sshiftr_clamp: assumes "LENGTH('a) \ n" shows "(x::'a::len word) >>> n = x >>> LENGTH('a)" apply (cases "0 \ sint x") subgoal apply (subst sshiftr_clamp_pos[OF assms]) defer apply (subst sshiftr_clamp_pos) by auto apply (subst sshiftr_clamp_neg[OF assms]) defer apply (subst sshiftr_clamp_neg) by auto text\ Like @{thm shiftr1_bl_of}, but the precondition is stronger because we need to pick the msb out of the list. \ lemma sshiftr1_bl_of: "length bl = LENGTH('a) \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd bl # butlast bl)" apply (rule word_bl.Rep_eqD) apply (subst bl_sshiftr1[of "of_bl bl :: 'a word"]) by (simp add: word_bl.Abs_inverse) text\ Like @{thm sshiftr1_bl_of}, with a weaker precondition. We still get a direct equation for @{term \sshiftr1 (of_bl bl)\}, it's just uglier. \ lemma sshiftr1_bl_of': "LENGTH('a) \ length bl \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd (drop (length bl - LENGTH('a)) bl) # butlast (drop (length bl - LENGTH('a)) bl))" apply (subst of_bl_drop'[symmetric, of "length bl - LENGTH('a)"]) using sshiftr1_bl_of[of "drop (length bl - LENGTH('a)) bl"] by auto text\ Like @{thm shiftr_bl_of}. \ lemma sshiftr_bl_of: assumes "length bl = LENGTH('a)" shows "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof - { fix n assume "n \ LENGTH('a)" hence "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof (induction n) case (Suc n) hence "n < length bl" by (simp add: assms) hence ne: "\take (length bl - n) bl = []" by auto have left: "hd (replicate n (hd bl) @ take (length bl - n) bl) = (hd bl)" by (cases "0 < n") auto have right: "butlast (take (length bl - n) bl) = take (length bl - Suc n) bl" by (subst butlast_take) auto have "(of_bl bl::'a::len word) >>> Suc n = sshiftr1 ((of_bl bl::'a::len word) >>> n)" unfolding sshiftr_eq_funpow_sshiftr1 by simp also have "\ = of_bl (replicate (Suc n) (hd bl) @ take (length bl - Suc n) bl)" apply (subst Suc.IH[OF Suc_leD[OF Suc.prems]]) apply (subst sshiftr1_bl_of) subgoal using assms Suc.prems by simp apply (rule arg_cong[where f=of_bl]) apply (subst butlast_append) unfolding left right using ne by simp finally show ?case . qed (transfer, simp) } note pos = this { assume n: "LENGTH('a) \ n" have "(of_bl bl::'a::len word) >>> n = (of_bl bl::'a::len word) >>> LENGTH('a)" by (rule sshiftr_clamp[OF n]) also have "\ = of_bl (replicate LENGTH('a) (hd bl) @ take (length bl - LENGTH('a)) bl)" apply (rule pos) .. also have "\ = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof - have "(of_bl (replicate LENGTH('a) (hd bl)) :: 'a word) = of_bl (replicate n (hd bl))" apply (subst of_bl_drop'[symmetric, of "n - LENGTH('a)" "replicate n (hd bl)"]) unfolding length_replicate by (auto simp: n) thus ?thesis by (simp add: assms n) qed finally have "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" . } thus ?thesis using pos by fastforce qed text\Like @{thm shiftr_bl}\ lemma sshiftr_bl: "x >>> n \ of_bl (replicate n (msb x) @ take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" unfolding word_msb_alt by (smt (z3) length_to_bl_eq sshiftr_bl_of word_bl.Rep_inverse) end diff --git a/thys/Word_Lib/Traditional_Infix_Syntax.thy b/thys/Word_Lib/Traditional_Infix_Syntax.thy --- a/thys/Word_Lib/Traditional_Infix_Syntax.thy +++ b/thys/Word_Lib/Traditional_Infix_Syntax.thy @@ -1,1077 +1,1077 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Operation variants with traditional syntax\ theory Traditional_Infix_Syntax imports "HOL-Library.Word" More_Word Signed_Words begin class semiring_bit_syntax = semiring_bit_shifts begin definition test_bit :: \'a \ nat \ bool\ (infixl "!!" 100) where test_bit_eq_bit: \test_bit = bit\ definition shiftl :: \'a \ nat \ 'a\ (infixl "<<" 55) where shiftl_eq_push_bit: \a << n = push_bit n a\ definition shiftr :: \'a \ nat \ 'a\ (infixl ">>" 55) where shiftr_eq_drop_bit: \a >> n = drop_bit n a\ end instance word :: (len) semiring_bit_syntax .. context includes lifting_syntax begin lemma test_bit_word_transfer [transfer_rule]: \(pcr_word ===> (=)) (\k n. n < LENGTH('a) \ bit k n) (test_bit :: 'a::len word \ _)\ by (unfold test_bit_eq_bit) transfer_prover lemma shiftl_word_transfer [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (\k n. push_bit n k) shiftl\ by (unfold shiftl_eq_push_bit) transfer_prover lemma shiftr_word_transfer [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (\k n. (drop_bit n \ take_bit LENGTH('a)) k) (shiftr :: 'a::len word \ _)\ by (unfold shiftr_eq_drop_bit) transfer_prover end lemma test_bit_word_eq: \test_bit = (bit :: 'a::len word \ _)\ by (fact test_bit_eq_bit) lemma shiftl_word_eq: \w << n = push_bit n w\ for w :: \'a::len word\ by (fact shiftl_eq_push_bit) lemma shiftr_word_eq: \w >> n = drop_bit n w\ for w :: \'a::len word\ by (fact shiftr_eq_drop_bit) lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" for u v :: "'a::len word" by (simp add: bit_eq_iff test_bit_eq_bit fun_eq_iff) lemma test_bit_size: "w !! n \ n < size w" for w :: "'a::len word" by transfer simp lemma word_eq_iff: "x = y \ (\n?P \ ?Q\) for x y :: "'a::len word" by transfer (auto simp add: bit_eq_iff bit_take_bit_iff) lemma word_eqI: "(\n. n < size u \ u !! n = v !! n) \ u = v" for u :: "'a::len word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ u !! x = v !! x" for u v :: "'a::len word" by simp lemma test_bit_bin': "w !! n \ n < size w \ bit (uint w) n" by transfer (simp add: bit_take_bit_iff) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma word_test_bit_def: \test_bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemma word_test_bit_transfer [transfer_rule]: "(rel_fun pcr_word (rel_fun (=) (=))) (\x n. n < LENGTH('a) \ bit x n) (test_bit :: 'a::len word \ _)" by (simp only: test_bit_eq_bit) transfer_prover lemma test_bit_wi [simp]: "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bit x n" by transfer simp lemma word_ops_nth_size: "n < size x \ (x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n) \ (x XOR y) !! n = (x !! n \ y !! n) \ (NOT x) !! n = (\ x !! n)" for x :: "'a::len word" by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff) lemma word_ao_nth: "(x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n)" for x :: "'a::len word" by transfer (auto simp add: bit_or_iff bit_and_iff) lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] lemmas msb1 = msb0 [where i = 0] lemma test_bit_numeral [simp]: "(numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bit (numeral w :: int) n" by transfer (rule refl) lemma test_bit_neg_numeral [simp]: "(- numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bit (- numeral w :: int) n" by transfer (rule refl) lemma test_bit_1 [iff]: "(1 :: 'a::len word) !! n \ n = 0" by transfer (auto simp add: bit_1_iff) lemma nth_0 [simp]: "\ (0 :: 'a::len word) !! n" by transfer simp lemma nth_minus1 [simp]: "(-1 :: 'a::len word) !! n \ n < LENGTH('a)" by transfer simp lemma shiftl1_code [code]: \shiftl1 w = push_bit 1 w\ by transfer (simp add: ac_simps) lemma uint_shiftr_eq: \uint (w >> n) = uint w div 2 ^ n\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) lemma shiftr1_code [code]: \shiftr1 w = drop_bit 1 w\ by transfer (simp add: drop_bit_Suc) lemma shiftl_def: \w << n = (shiftl1 ^^ n) w\ proof - have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) then show ?thesis by transfer simp qed lemma shiftr_def: \w >> n = (shiftr1 ^^ n) w\ proof - have \shiftr1 ^^ n = (drop_bit n :: 'a word \ 'a word)\ apply (induction n) apply simp apply (simp only: shiftr1_eq_div_2 [abs_def] drop_bit_eq_div [abs_def] funpow_Suc_right) apply (use div_exp_eq [of _ 1, where ?'a = \'a word\] in simp) done then show ?thesis by (simp add: shiftr_eq_drop_bit) qed lemma bit_shiftl_word_iff [bit_simps]: \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ for w :: \'a::len word\ by (simp add: shiftl_word_eq bit_push_bit_iff not_le) lemma bit_shiftr_word_iff [bit_simps]: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: shiftr_word_eq bit_drop_bit_eq) lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - Suc 0) k))\ by (simp flip: signed_take_bit_decr_length_iff) lemma sshiftr_eq [code]: \w >>> n = signed_drop_bit n w\ by transfer simp lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (sshiftr1 ^^ n) w\ apply (rule sym) apply (simp add: sshiftr1_eq_signed_drop_bit_Suc_0 sshiftr_eq) apply (induction n) apply simp_all done lemma uint_sshiftr_eq: \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ for w :: \'a::len word\ by transfer (simp flip: drop_bit_eq_div) lemma sshift1_code [code]: \sshiftr1 w = signed_drop_bit 1 w\ by transfer (simp add: drop_bit_Suc) lemma sshiftr_0 [simp]: "0 >>> n = 0" by transfer simp lemma sshiftr_n1 [simp]: "-1 >>> n = -1" by transfer simp lemma bit_sshiftr_word_iff [bit_simps]: \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ for w :: \'a::len word\ apply transfer apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma nth_sshiftr : "(w >>> m) !! n = (n < size w \ (if n + m \ size w then w !! (size w - 1) else w !! (n + m)))" apply transfer apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le ac_simps) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma sshiftr_numeral [simp]: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ apply (rule word_eqI) apply (cases \LENGTH('a)\) apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr bit_signed_take_bit_iff min_def not_le not_less less_Suc_eq_le ac_simps) done setup \ Context.theory_map (fold SMT_Word.add_word_shift' [ (\<^term>\shiftl :: 'a::len word \ _\, "bvshl"), (\<^term>\shiftr :: 'a::len word \ _\, "bvlshr"), (\<^term>\sshiftr :: 'a::len word \ _\, "bvashr") ]) \ lemma revcast_down_us [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >>> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_ucast_iff bit_sshiftr_word_iff ac_simps) done lemma revcast_down_ss [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >>> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_scast_iff bit_sshiftr_word_iff ac_simps) done lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" using sint_signed_drop_bit_eq [of n w] by (simp add: drop_bit_eq_div sshiftr_eq) lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" shows "bit (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def) lemma test_bit_2p: "(word_of_int (2 ^ n)::'a::len word) !! m \ m = n \ m < LENGTH('a)" by transfer (auto simp add: bit_exp_iff) lemma nth_w2p: "((2::'a::len word) ^ n) !! m \ m = n \ m < LENGTH('a::len)" by transfer (auto simp add: bit_exp_iff) lemma bang_is_le: "x !! m \ 2 ^ m \ x" for x :: "'a::len word" apply (rule xtrans(3)) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ by transfer (simp add: mask_eq_exp_minus_1 push_bit_of_1) lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) lemma shiftl_0 [simp]: "(0::'a::len word) << n = 0" by transfer simp lemma shiftr_0 [simp]: "(0::'a::len word) >> n = 0" by transfer simp lemma nth_shiftl1: "shiftl1 w !! n \ n < size w \ n > 0 \ w !! (n - 1)" by transfer (auto simp add: bit_double_iff) lemma nth_shiftl': "(w << m) !! n \ n < size w \ n >= m \ w !! (n - m)" for w :: "'a::len word" by transfer (auto simp add: bit_push_bit_iff) lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" by transfer (auto simp add: bit_take_bit_iff simp flip: bit_Suc) lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct "m" arbitrary: n) apply (auto simp add: nth_shiftr1) done lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" apply transfer apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" by (fact uint_shiftr_eq) lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) lemma shiftl_numeral [simp]: \numeral k << numeral l = (push_bit (numeral l) (numeral k) :: 'a::len word)\ by (fact shiftl_word_eq) lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len word" apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (induct n) (auto simp: shiftl_def shiftl1_2t) lemma shiftr_numeral [simp]: \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ by (fact shiftr_word_eq) lemma shiftr_numeral_Suc [simp]: \(numeral k >> Suc 0 :: 'a::len word) = drop_bit (Suc 0) (numeral k)\ by (fact shiftr_word_eq) lemma drop_bit_numeral_bit0_1 [simp]: \drop_bit (Suc 0) (numeral k) = (word_of_int (drop_bit (Suc 0) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by (metis Word_eq_word_of_int drop_bit_word.abs_eq of_int_numeral) lemma nth_mask [simp]: \(mask n :: 'a::len word) !! i \ i < n \ i < size (mask n :: 'a word)\ by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) lemma slice_shiftr: "slice n w = ucast (w >> n)" apply (rule bit_word_eqI) apply (cases \n \ LENGTH('b)\) apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps dest: bit_imp_le_length) done lemma nth_slice: "(slice n w :: 'a::len word) !! m = (w !! (m + n) \ m < LENGTH('a))" by (simp add: slice_shiftr nth_ucast nth_shiftr) lemma revcast_down_uu [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_ucast_iff bit_shiftr_word_iff ac_simps) done lemma revcast_down_su [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_scast_iff bit_shiftr_word_iff ac_simps) done lemma cast_down_rev [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (w << n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) done lemma revcast_up [OF refl]: "rc = revcast \ source_size rc + n = target_size rc \ rc w = (ucast w :: 'a::len word) << n" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) apply auto apply (metis add.commute add_diff_cancel_right) apply (metis diff_add_inverse2 diff_diff_add) done lemmas rc1 = revcast_up [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas rc2 = revcast_down_uu [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas ucast_up = rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] lemmas ucast_down = rc2 [simplified rev_shiftr revcast_ucast [symmetric]] \ \problem posed by TPHOLs referee: criterion for overflow of addition of signed integers\ lemma sofl_test: \sint x + sint y = sint (x + y) \ (x + y XOR x) AND (x + y XOR y) >> (size x - 1) = 0\ for x y :: \'a::len word\ proof - obtain n where n: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] by (auto intro: ccontr) have \sint x + sint y = sint (x + y) \ (sint (x + y) < 0 \ sint x < 0) \ (sint (x + y) < 0 \ sint y < 0)\ using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] apply (auto simp add: not_less) apply (unfold sint_word_ariths) apply (subst signed_take_bit_int_eq_self) prefer 4 apply (subst signed_take_bit_int_eq_self) prefer 7 apply (subst signed_take_bit_int_eq_self) prefer 10 apply (subst signed_take_bit_int_eq_self) apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) apply (smt (z3) take_bit_nonnegative) apply (smt (z3) take_bit_int_less_exp) apply (smt (z3) take_bit_nonnegative) apply (smt (z3) take_bit_int_less_exp) done then show ?thesis apply (simp only: One_nat_def word_size shiftr_word_eq drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) apply (simp add: bit_last_iff) done qed lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) lemma test_bit_cat [OF refl]: "wc = word_cat a b \ wc !! n = (n < size wc \ (if n < size b then b !! n else a !! (n - size b)))" apply (simp add: word_size not_less; transfer) apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) done \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. b !! n = (n < size b \ c !! n) \ a !! m = (m < size a \ c !! (m + size b)))" by (auto simp add: word_split_bin' test_bit_bin bit_unsigned_iff word_size bit_drop_bit_eq ac_simps dest: bit_imp_le_length) lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. b !! n \ n < size b \ c !! n) \ (\m::nat. a !! m \ m < size a \ c !! (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. b !! n = (n < size b \ c !! n)) \ (\m::nat. a !! m = (m < size a \ c !! (m + size b))))" apply (rule_tac iffI) apply (rule_tac conjI) apply (erule test_bit_split [THEN conjunct1]) apply (erule test_bit_split [THEN conjunct2]) apply (case_tac "word_split c") apply (frule test_bit_split) apply (erule trans) apply (fastforce intro!: word_eqI simp add: word_size) done lemma test_bit_rcat: "sw = size (hd wl) \ rc = word_rcat wl \ rc !! n = (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" for wl :: "'a::len word list" by (simp add: word_size word_rcat_def foldl_map rev_map bit_horner_sum_uint_exp_iff) (simp add: test_bit_eq_bit) lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] -lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" +lemma max_test_bit: "(- 1::'a::len word) !! n \ n < LENGTH('a)" by (fact nth_minus1) lemma shiftr_x_0 [iff]: "x >> 0 = x" for x :: "'a::len word" by transfer simp lemma shiftl_x_0 [simp]: "x << 0 = x" for x :: "'a::len word" by (simp add: shiftl_t2n) lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" by (simp add: shiftl_t2n) lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" by (induct xs) auto lemma word_and_1: "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff test_bit_eq_bit bit_1_iff intro: gr0I) lemma test_bit_1' [simp]: "(1 :: 'a :: len word) !! n \ 0 < LENGTH('a) \ n = 0" by simp lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) lemma word_ops_nth [simp]: fixes x y :: \'a::len word\ shows word_or_nth: "(x OR y) !! n = (x !! n \ y !! n)" and word_and_nth: "(x AND y) !! n = (x !! n \ y !! n)" and word_xor_nth: "(x XOR y) !! n = (x !! n \ y !! n)" by ((cases "n < size x", auto dest: test_bit_size simp: word_ops_nth_size word_size)[1])+ lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma nth_w2p_same: "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" by (simp add : nth_w2p) lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add: uint_2p_alt word_size) apply (metis uint_shiftr_eq word_of_int_uint) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply (unfold shiftr_def) apply (induct_tac "n") apply auto apply (erule le_shiftr1) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" apply (rule word_eqI) apply (simp add: word_size nth_shiftr) done lemma shiftr_mask [simp]: \mask m >> m = (0::'a::len word)\ by (rule shiftr_mask_le) simp lemma word_leI: "(\n. \n < size (u::'a::len word); u !! n \ \ (v::'a::len word) !! n) \ u <= v" apply (rule xtrans(4)) apply (rule word_and_le2) apply (rule word_eqI) apply (simp add: word_ao_nth) apply safe apply assumption apply (erule_tac [2] asm_rl) apply (unfold word_size) by auto lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" for w :: \'a::len word\ apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtrans(3)) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" for w :: \'a::len word\ apply (unfold test_bit_eq_iff [THEN sym]) apply (rule iffI) apply (rule ext) apply (rule_tac [2] ext) apply (auto simp add : word_ao_nth nth_shiftr) apply (drule arg_cong) apply (drule iffD2) apply assumption apply (simp add : word_ao_nth) prefer 2 apply (simp add : word_size test_bit_bin) apply transfer apply (auto simp add: fun_eq_iff bit_simps) apply (metis add_diff_inverse_nat) done lemma mask_shiftl_decompose: "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" by (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. x !! n = y !! n)" by (subst test_bit_eq_iff[symmetric]) fastforce lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" apply(rule word_eqI) apply(simp add: word_ao_nth nth_shiftl, safe) done lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" apply(rule word_eqI) apply(simp add:nth_shiftl word_ao_nth, safe) done lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply arith done lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma mask_shift: "(x AND NOT (mask y)) >> y = x >> y" for x :: \'a::len word\ apply (rule bit_eqI) apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) using bit_imp_le_length apply auto done lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_eq_nat_uint) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_eq_unatI) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by transfer (simp add: take_bit_push_bit) lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') apply transfer apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) done lemma neg_mask_test_bit: "(NOT(mask n) :: 'a :: len word) !! m = (n \ m \ m < LENGTH('a))" by (metis not_le nth_mask test_bit_bin word_ops_nth_size word_size) lemma upper_bits_unset_is_l2p: \(\n' \ n. n' < LENGTH('a) \ \ p !! n') \ (p < 2 ^ n)\ (is \?P \ ?Q\) if \n < LENGTH('a)\ for p :: "'a :: len word" proof assume ?Q then show ?P by (meson bang_is_le le_less_trans not_le word_power_increasing) next assume ?P have \take_bit n p = p\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ show \bit (take_bit n p) q \ bit p q\ proof (cases \q < n\) case True then show ?thesis by (auto simp add: bit_simps) next case False then have \n \ q\ by simp with \?P\ \q < LENGTH('a)\ have \\ bit p q\ by (simp add: test_bit_eq_bit) then show ?thesis by (simp add: bit_simps) qed qed with that show ?Q using take_bit_word_eq_self_iff [of n p] by auto qed lemma less_2p_is_upper_bits_unset: "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n')" for p :: "'a :: len word" by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) lemma test_bit_over: "n \ size (x::'a::len word) \ (x !! n) = False" by transfer auto lemma le_mask_high_bits: "w \ mask n \ (\i \ {n ..< size w}. \ w !! i)" for w :: \'a::len word\ by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) lemma test_bit_conj_lt: "(x !! m \ m < LENGTH('a)) = x !! m" for x :: "'a :: len word" using test_bit_bin by blast lemma neg_test_bit: "(NOT x) !! n = (\ x !! n \ n < LENGTH('a))" for x :: "'a::len word" by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) lemma shiftr_less_t2n': "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit ac_simps) done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" by (rule bit_word_eqI) (auto simp add: bit_simps dest: bit_imp_le_length) lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" by transfer (auto simp add: bit_simps) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done lemma nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (rule ccontr) apply (auto simp add: not_less test_bit_word_eq) apply (meson bit_imp_le_length bit_uint_iff less_2p_is_upper_bits_unset test_bit_bin) done lemma shiftl_mask_is_0[simp]: "(x << n) AND mask n = 0" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask add: shiftl_eq_push_bit take_bit_push_bit) lemma rshift_sub_mask_eq: "(a >> (size a - b)) AND mask b = a >> (size a - b)" for a :: \'a::len word\ using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_eq_decr_exp word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" for a :: \'a::len word\ apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" for w :: \'a::len word\ by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" for w :: \'a::len word\ by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" for x :: \'a::len word\ by (simp add: le_mask_iff shiftl_shiftr3) lemma word_and_1_shiftl: "x AND (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x AND (mask n << m) = ((x >> m) AND mask n) << m" for x :: \'a::len word\ apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma of_bool_nth: "of_bool (x !! v) = (x >> v) AND 1" for x :: \'a::len word\ by (simp add: test_bit_word_eq shiftr_word_eq bit_eq_iff) (auto simp add: bit_1_iff bit_and_iff bit_drop_bit_eq intro: ccontr) lemma shiftr_mask_eq: "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit) done lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x AND y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x AND (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply (rule bit_word_eqI) apply (auto simp add: bit_simps test_bit_eq_bit) done lemma and_neq_0_is_nth: \x AND y \ 0 \ x !! n\ if \y = 2 ^ n\ for x y :: \'a::len word\ apply (simp add: bit_eq_iff bit_simps) using that apply (simp add: bit_simps not_le) apply transfer apply auto done lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x AND 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" for w :: \'a::len word\ by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff shiftl_word_eq bit_push_bit_iff) end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,1901 +1,1901 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Type_Syntax Signed_Division_Word Signed_Words More_Word Most_significant_bit Enumeration_Word Aligned begin +lemma max_word_not_less [simp]: + "\ - 1 < x" for x :: \'a::len word\ + by (fact word_order.extremum_strict) + (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" for x :: \'a::len word\ by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" for a b :: \'a::len word\ apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_eq_decr_exp) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff min_def) lemma bit_twiddle_max: "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma swap_with_xor: "\(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y\ \ z = b \ y = a" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma scast_nop1 [simp]: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop2 [simp]: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x AND mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)" by (auto simp add: Parity.bit_eq_iff bit_not_iff bit_or_iff bit_and_iff bit_mask_iff) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w AND NOT(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) AND mask m) AND mask n = x AND mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by transfer (simp add: drop_bit_Suc) lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (cases \n = 0\) apply clarsimp apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" using even_plus_one_iff [of x] by (simp add: test_bit_word_eq) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" by transfer (simp add: even_nat_iff) lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" apply (cases \LENGTH('a)\; transfer) apply (simp_all add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit drop_bit_Suc) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) OR (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma mask_or_not_mask: "x AND mask n OR x AND NOT (mask n) = x" for x :: \'a::len word\ apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) AND NOT (mask n) = p AND NOT (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0..UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by simp have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis apply (simp add: word_ubin.eq_norm bintrunc_mod2p unsigned_ucast_eq) apply (metis \x \ 2 ^ LENGTH('b) - 1\ and_mask_eq_iff_le_mask mask_eq_decr_exp take_bit_eq_mask) done qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) lemma complement_nth_w2p: shows "n' < LENGTH('a) \ (NOT (2 ^ n :: 'a::len word)) !! n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x AND y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) lemma le_or_mask: "w \ w' \ w OR mask x \ w' OR mask x" for w w' :: \'a::len word\ by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply transfer apply simp done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (induct n; simp add: shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) apply (fastforce dest: le_shiftr1') done -lemma word_add_no_overflow:"(x::'a::len word) < max_word \ x < x + 1" +lemma word_add_no_overflow:"(x::'a::len word) < - 1 \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" unfolding ucast_eq unat_eq_nat_uint apply (subst int_word_uint) apply (subst mod_pos_pos_trivial; simp?) apply (rule lt2p_lem) apply (simp add: assms) done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" using unat_less_helper unat_ucast_less_no_overflow by blast lemma unat_ucast_no_overflow_le: assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) apply (simp only: flip: ucast_nat_def) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" by (simp add: and_not_mask shiftr_eq_0) lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) shiftl_shiftr_id shiftl_t2n) lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_diff_power_eq) done lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (meson nat_mult_power_less_eq zero_less_numeral) done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) apply transfer apply (auto simp add: is_aligned_iff_udvd) apply (metis bintrunc_bintrunc_ge int_ops(1) nat_int_comparison(1) nat_less_le take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat) done lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt2 lt3 unat_mono word_less_minus_mono_left by blast qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: - "(max_word :: 'a::len word) = mask LENGTH('a)" - unfolding mask_eq_decr_exp by simp + "(- 1 :: 'a::len word) = mask LENGTH('a)" + by (fact minus_1_eq_mask) lemmas mask_len_max = max_word_mask[symmetric] lemma mask_out_first_mask_some: "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" for x y :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma mask_lower_twice: "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma mask_lower_twice2: "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" for a :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_and_neg_mask: "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_and_mask: "ucast (x AND mask n) = ucast x AND mask n" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done (* negating a mask which has been shifted to the very left *) lemma NOT_mask_shifted_lenword: "NOT (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)" by (rule bit_word_eqI) (auto simp add: shiftl_word_eq word_size bit_not_iff bit_push_bit_iff bit_mask_iff) (* Comparisons between different word sizes. *) lemma eq_ucast_ucast_eq: "LENGTH('b) \ LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by transfer simp lemma le_ucast_ucast_le: "x \ ucast y \ ucast x \ y" for x :: "'a::len word" and y :: "'b::len word" by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) lemma less_ucast_ucast_less: "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" for x :: "'a::len word" and y :: "'b::len word" by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) lemma ucast_le_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" for x :: "'a::len word" by (simp add: unat_arith_simps(1) unat_ucast_up_simp) lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" by transfer simp lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w AND w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" by (simp add: bit_eq_iff test_bit_word_eq) lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" using unat_gt_0 [of \- 1 :: 'a::len word\] by simp (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" shows "x = y" proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ from eq have \push_bit n x = push_bit n y\ by (simp add: push_bit_eq_mult) moreover from le have \take_bit m x = x\ \take_bit m y = y\ by (simp_all add: less_eq_mask_iff_take_bit_eq_self) ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ by simp_all with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ apply (simp add: push_bit_take_bit) unfolding bit_eq_iff apply (simp add: bit_simps not_le) apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) apply (meson less_irrefl not_le zero_le_numeral zero_le_power) done lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num) lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: \bin_cat a n b \ uints m\ if \a \ uints l\ \m \ l + n\ proof - from \m \ l + n\ obtain q where \m = l + n + q\ using le_Suc_ex by blast then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ by (simp add: ac_simps power_add) moreover have \a mod 2 ^ (l + q) = a\ using \a \ uints l\ by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) then show ?thesis by (simp add: uints_def) qed lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" using word_of_int_bin_cat_eq_iff [OF that, of b a d c] by transfer auto lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = (2 :: 'a::len word) ^ n + mask n" by (simp add: mask_eq_decr_exp) lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a AND b = 0" by (metis and_zero_eq is_aligned_mask le_mask_imp_and_mask word_bw_lcs(1)) lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a OR b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a OR b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" apply transfer apply (auto simp add: min_def) apply (metis bintrunc_bintrunc_ge bintrunc_n_0 nat_less_le not_le take_bit_eq_0_iff) done lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" apply (simp add: less_eq_mask_iff_take_bit_eq_self) apply transfer apply (simp add: ac_simps) done lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: "\ p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ \ p' = p \ off' = off" apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] ucast_leq_mask) apply (simp add: is_aligned_nth) apply (rule conjI; word_eqI) apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) apply (rename_tac i) apply (erule_tac x="i+n" in allE) apply simp done lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed -lemma max_word_not_less[simp]: - "\ max_word < x" - by (simp add: not_less) - lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: \(x + y * m) AND NOT(mask n) = (x AND NOT(mask n)) + y * m\ if \m AND (2 ^ n - 1) = 0\ for x y m :: \'a::len word\ by (metis (no_types, hide_lams) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: of_nat_diff unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y AND mask n = 0 \ x + y AND NOT(mask n) = (x AND NOT(mask n)) + y" for x y :: \'a::len word\ by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x AND mask n = 0; y AND mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_iff_dvd_nat) done lemma sub_right_shift: "\ x AND mask n = 0; y AND mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y AND mask n = mask n \ (x AND y) AND mask n = x AND mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y AND mask n = 0 \ (x AND y) AND mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) AND b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) AND b < c" by transfer simp lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma unsigned_uminus1 [simp]: \(unsigned (-1::'b::len word)::'c::len word) = mask LENGTH('b)\ by word_eqI lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" by word_eqI_solve lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x AND mask n = 0; m \ n \ \ x AND mask m = 0" for x :: \'a::len word\ by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) AND mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) AND 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma shiftr_and_eq_shiftl: "(w >> n) AND x = y \ w AND (x << n) = (y << n)" for y :: "'a:: len word" by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" apply (simp add: word_le_def) apply (simp only: uint_nat zle_int) apply transfer apply (simp add: take_bit_nat_eq_self) done lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x AND NOT(mask sz)) + 2 ^ sz - 1" for x :: \'a::len word\ by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" by transfer (simp add: take_bit_add) lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) AND (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) AND mask m) = unat (x AND mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_eq_decr_exp power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit) apply (rule order_trans) defer apply assumption apply (simp add: nat_le_iff of_nat_diff) done lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n AND msk = mask n \ \ (x mod m) AND msk = x mod m" for x :: \'a::len word\ by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x AND mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m AND NOT(mask n :: 'a::len word)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr AND NOT(mask sz')) - (ptr AND NOT(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") for ptr :: \'a::len word\ proof - assume lt: "sz' \ sz" hence "?lhs = ptr AND (mask sz AND NOT(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) AND NOT(mask sz) = 0" by (simp add: of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr AND NOT(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) AND NOT(mask sz)) + unat (ptr AND mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" by (metis (no_types) of_nat_power diff_le_self le_less_trans shiftl_less_t2n unat_less_power word_unat.Rep_inverse) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d AND mask n) = unat d \ unat (p + d AND NOT(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) AND NOT(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) AND mask high_bits = x >> low_bits" for x :: \'a::len word\ by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x AND mask low_bits = 0) = (x = 0)" for x :: \'a::len word\ using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) AND mask n = q AND mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) AND mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x AND NOT(mask sz)) + (x AND mask sz) \ (x AND NOT(mask sz)) + ((x AND mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "NOT(mask n :: 'a::len word) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * NOT(mask n) = - (x AND NOT(mask n))" for x :: \'a::len word\ by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) AND NOT(mask n) = p + (x AND NOT(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) AND NOT(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) AND mask (m::nat) = (mask (min m n) :: 'a::len word)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n AND mask m = (if n < m then 2 ^ n else (0 :: 'a::len word))" by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: - "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (max_word :: 'a :: len word) \ + "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (- 1 :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p AND mask a >> b) :: 'a :: len word) = p AND mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p AND mask a >> b) :: 'a :: len word) = unat (p AND mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p AND mask a) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p AND mask a << c) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p AND NOT(mask a)) AND mask b = 0" for p :: \'a::len word\ using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p AND NOT(mask a)) + ((p AND mask a) AND NOT(mask b)) + (p AND mask b) = p" for p :: \'a::len word\ by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p AND mask a >> b << b) = (p AND mask a) AND NOT(mask b)" for p :: \'a::len word\ by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p AND mask b) \ \ (p AND NOT(mask a)) + (p AND mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (fact Word.of_int_sint) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by transfer simp lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x AND mask n \ of_nat ` set [0 ..< 2 ^ n]" apply (simp flip: take_bit_eq_mask) apply (rule image_eqI [of _ _ \unat (take_bit n x)\]) using len apply simp_all apply transfer apply simp done lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w AND mask (size w - n)" for w :: \'a::len word\ by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp add: mask_eq_decr_exp NOT_eq flip: mul_not_mask_eq_neg_shiftl) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x AND mask LENGTH('b) = y AND mask LENGTH('b))" by (rule iffI; word_eqI_solve) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" by (metis assms sbintrunc_bintrunc ucast_eq word_ubin.eq_norm) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i = (if n < i then w !! n else w !! i)" using assms by (simp add: nth_sbintr) (simp add: test_bit_bin) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" apply (subst sbintrunc_uint_ucast) apply simp apply (subst test_bit_sbintrunc) apply (rule len_a) apply (rule if_cong[OF _ refl refl]) using leD less_linear by fastforce lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ ultimately have \LENGTH('b) = m + LENGTH('a)\ by simp then show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (rule bit_word_eqI) apply (simp add: bit_signed_take_bit_iff) done next case False define q where \q = LENGTH('b) - 1\ then have \LENGTH('b) = Suc q\ by simp moreover define m where \m = Suc LENGTH('a) - LENGTH('b)\ with False \LENGTH('b) = Suc q\ have \LENGTH('a) = m + q\ by (simp add: not_le) ultimately show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (transfer fixing: m q) apply (simp add: signed_take_bit_take_bit) apply rule apply (subst bit_eq_iff) apply (simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def) apply (auto simp add: Suc_le_eq) using less_imp_le_nat apply blast using less_imp_le_nat apply blast done qed lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ NOT(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x AND mask LENGTH('b))" by word_eqI_solve lemma ucast_NOT: "ucast (NOT x) = NOT(ucast x) AND mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (NOT x) = NOT(UCAST('a \ 'b) x)" by word_eqI lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_ubin.norm_eq_iff [symmetric]) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_exp_eq min_def) apply (rule distrib [symmetric]) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_eq sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (simp only: scast_eq) apply (metis down_cast_same is_up_down scast_eq ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: signed_divide_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ end