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,1030 +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" 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" 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 auto - subgoal by transfer (auto simp: emax_def unat_sub) - subgoal - apply transfer - apply safe - by (metis One_nat_def diff_less lessI minus_one_norm minus_one_word unat_0 unat_lt2p - unat_of_nat_len) - subgoal by transfer auto - subgoal by transfer (auto simp: emax_def unat_sub) - subgoal - apply transfer - apply safe - by (metis One_nat_def diff_less lessI minus_one_norm minus_one_word unat_0 unat_lt2p - unat_of_nat_len) + subgoal by transfer simp + subgoal by transfer (simp add: emax_eq bintrunc_eq_take_bit 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 bintrunc_eq_take_bit 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 -lemma emax_eq: "emax x = 2^LENGTH('e) - 1" - for x::"('e, 'f)float itself" - unfolding bias_def emax_def - by (auto simp: max_word_eq unat_minus word_size) - 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) 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 ) 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: 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) 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 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) 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,412 +1,412 @@ (* Title: IP_Address.thy Authors: Cornelius Diekmann *) theory IP_Address imports "Word_Lib.Word_Lemmas" 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" - by(simp add: max_ip_addr_def max_word_def word_of_int_minus) + by(simp add: max_ip_addr_def word_of_int_minus) 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) 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 Let_def mask_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 Word.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.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,269 +1,269 @@ (* 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 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" - by(simp add: max_ipv4_addr_number max_word_def) + 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: "\ n \ nat_of_ipv4addr max_ipv4_addr \ \ nat_of_ipv4addr (ipv4addr_of_nat n) = n" by (simp add: ipv4addr_of_nat_def le_unat_uoi nat_of_ipv4addr_def) 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) 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 word_of_nat) apply(simp add: nat_of_ipv4addr_def unat_def) apply(simp add: and_mask_mod_2p) apply(simp add: shiftr_div_2n) apply(simp add: uint_word_of_int) 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 and_mask_mod_2p) apply(simp add: word_of_nat) (*use this to get rid of of_nat. All thm are with word_of_int*) apply(simp add: uint_word_of_int) apply(subst mod_mod_cancel) apply(simp; fail) apply(simp add: zmod_int) 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 word_of_nat) apply(simp add: nat_of_ipv4addr_def unat_def) apply(simp add: and_mask_mod_2p) apply(simp add: shiftr_div_2n) apply(simp add: uint_word_of_int) apply(simp add: NumberWang_IPv4.div65536[simplified]) (*The [simplified] is needed because Word_Lib adds some additional simp rules*) done \ \When @{file \../Word_Lib/Word_Lemmas.thy\} is imported, some @{file \NumberWang_IPv4.thy\} lemmas need the [simplified] attribute because @{text Word_Lib} adds some simp rules. This theory should also work without @{file \../Word_Lib/Word_Lemmas.thy\}\ 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 word_of_nat) apply(simp add: nat_of_ipv4addr_def unat_def) apply(simp add: and_mask_mod_2p) apply(simp add: shiftr_div_2n) apply(simp add: uint_word_of_int) apply(simp add: NumberWang_IPv4.div256[simplified]) 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) apply(simp add: ipv4addr_of_nat_def word_of_nat) apply(simp add: nat_of_ipv4addr_def) apply(subgoal_tac "(d + 256 * c + 65536 * b + 16777216 * a) mod 256 = d") apply(simp add: unat_def uint_word_of_int; fail) apply(simp add: NumberWang_IPv4.mod256) done from a b c d show ?thesis apply(simp add: ipv4addr_of_dotdecimal.simps dotdecimal_of_ipv4addr.simps) apply(simp add: mask_def) 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: Word_Lemmas.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_def) 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: Word.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: Word.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.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_def 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 mask_def) 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_dotdecimal.simps ipv4addr_of_nat_def ipset_from_cidr_def ipset_from_netmask_def mask_def) 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,953 +1,953 @@ (* 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 word_of_nat) 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" - by(simp add: max_ipv6_addr_number max_word_def) + 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: "n \ nat_of_ipv6addr max_ipv6_addr \ nat_of_ipv6addr (ipv6addr_of_nat n) = n" by(simp add: nat_of_ipv6addr_def ipv6addr_of_nat_def le_unat_uoi) 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) 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_def)+ 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 proof - (*sledgehammered for 128 word and concrete values for m and n*) have f1: "\w wa wb. ((w::'a::len word) && wa) && wb = w && wb && wa" by (simp add: word_bool_alg.conj_left_commute word_bw_comms(1)) have "\w n wa. ((w::'a::len word) && ~~ (mask n)) && (wa << n) = (w >> n) && wa << n" by (simp add: and_not_mask shiftl_over_and_dist) then show ?thesis by (simp add: is_aligned_mask is_aligned_shiftr_shiftl word_bool_alg.conj.assoc) (*using f1 by (metis (no_types) and_not_mask word_and_mask_shiftl word_bw_comms(1))*) qed 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_def) 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) (*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/Interval_Arithmetic_Word32/Interval_Word32.thy b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy --- a/thys/Interval_Arithmetic_Word32/Interval_Word32.thy +++ b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy @@ -1,3879 +1,3880 @@ (* Author: Brandon Bohrer *) theory Interval_Word32 imports Complex_Main Word_Lib.Word_Lemmas Word_Lib.Word_Lib Word_Lib.Word_Syntax begin text\Interval-Word32.thy implements conservative interval arithmetic operators on 32-bit word values, with explicit infinities for values outside the representable bounds. It is suitable for use in interpreters for languages which must have a well-understood low-level behavior (see Interpreter.thy). This work was originally part of the paper by Bohrer \emph{et al.}~\cite{BohrerTMMP18}. It is worth noting that this is not the first formalization of interval arithmetic in Isabelle/HOL. This article is presented regardless because it has unique goals in mind which have led to unique design decisions. Our goal is generate code which can be used to perform conservative arithmetic in implementations extracted from a proof. The Isabelle standard library now features interval arithmetic, for example in Approximation.thy. Ours differs in two ways: 1) We use intervals with explicit positive and negative infinities, and with overflow checking. Such checking is often relevant in implementation-level code with unknown inputs. To promote memory-efficient implementations, we moreover use sentinel values for infinities, rather than datatype constructors. This is especially important in real-time settings where the garbarge collection required for datatypes can be a concern. 2) Our goal is not to use interval arithmetic to discharge Isabelle goals, but to generate useful proven-correct implementation code, see Interpreter.thy. On the other hand, we are not concerned with producing interval-based automation for arithmetic goals in HOL. In practice, much of the work in this theory comes down to sheer case-analysis. Bounds-checking requires many edge cases in arithmetic functions, which come with many cases in proofs. Where possible, we attempt to offload interesting facts about word representations of numbers into reusable lemmas, but even then main results require many subcases, each with a certain amount of arithmetic grunt work. \ section \Interval arithmetic definitions\ subsection \Syntax\ text\Words are 32-bit\ type_synonym word = "32 Word.word" text\Sentinel values for infinities. Note that we leave the maximum value ($2^31$) completed unused, so that negation of $(2^{31})-1$ is not an edge case\ definition NEG_INF::"word" where NEG_INF_def[simp]:"NEG_INF = -((2 ^ 31) -1)" definition NegInf::"real" where NegInf[simp]:"NegInf = real_of_int (sint NEG_INF)" definition POS_INF::"word" where POS_INF_def[simp]:"POS_INF = (2^31) - 1" definition PosInf::"real" where PosInf[simp]:"PosInf = real_of_int (sint POS_INF)" text\Subtype of words who represent a finite value. \ typedef bword = "{n::word. sint n \ sint NEG_INF \ sint n \ sint POS_INF}" apply(rule exI[where x=NEG_INF]) by (auto) text\Numeric literals\ type_synonym lit = bword setup_lifting type_definition_bword lift_definition bword_zero::"bword" is "0::32 Word.word" by auto lift_definition bword_one::"bword" is "1::32 Word.word" by(auto simp add: sint_uint) lift_definition bword_neg_one::"bword" is "-1::32 Word.word" by(auto) definition word::"word \ bool" where word_def[simp]:"word w \ w \ {NEG_INF..POS_INF}" named_theorems rep_simps "Simplifications for representation functions" text\Definitions of interval containment and word representation repe w r iff word w encodes real number r\ inductive repe ::"word \ real \ bool" (infix "\\<^sub>E" 10) where repPOS_INF:"r \ real_of_int (sint POS_INF) \ repe POS_INF r" | repNEG_INF:"r \ real_of_int (sint NEG_INF) \ repe NEG_INF r" | repINT:"(sint w) < real_of_int(sint POS_INF) \ (sint w) > real_of_int(sint NEG_INF) \ repe w (sint w)" inductive_simps repePos_simps[rep_simps]:"repe POS_INF r" and repeNeg_simps[rep_simps]:"repe NEG_INF r" and repeInt_simps[rep_simps]:"repe w (sint w)" text\repU w r if w represents an upper bound of r\ definition repU ::"word \ real \ bool" (infix "\\<^sub>U" 10) where "repU w r \ \ r'. r' \ r \ repe w r'" lemma repU_leq:"repU w r \ r' \ r \ repU w r'" unfolding repU_def using order_trans by auto text\repU w r if w represents a lower bound of r\ definition repL ::"word \ real \ bool" (infix "\\<^sub>L" 10) where "repL w r \ \ r'. r' \ r \ repe w r'" lemma repL_geq:"repL w r \ r' \ r \ repL w r'" unfolding repL_def using order_trans by auto text\repP (l,u) r iff l and u encode lower and upper bounds of r\ definition repP ::"word * word \ real \ bool" (infix "\\<^sub>P" 10) where "repP w r \ let (w1, w2) = w in repL w1 r \ repU w2 r" lemma int_not_posinf: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ POS_INF" using b1 b2 by auto lemma int_not_neginf: assumes b1:" real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:" real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF" using b1 b2 by auto lemma int_not_undef: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF-1" using b1 b2 by auto lemma sint_range: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "sint ra \ {i. i > -((2^31)-1) \ i < (2^31)-1}" using b1 b2 by auto lemma word_size_neg: fixes w :: "32 Word.word" shows "size (-w) = size w" using Word.word_size[of w] Word.word_size[of "-w"] by auto lemma uint_distinct: fixes w1 w2 shows "w1 \ w2 \ uint w1 \ uint w2" by auto section \Preliminary lemmas\ subsection \Case analysis lemmas\ text\Case analysis principle for pairs of intervals, used in proofs of arithmetic operations\ lemma ivl_zero_case: fixes l1 u1 l2 u2 :: real assumes ivl1:"l1 \ u1" assumes ivl2:"l2 \ u2" shows "(l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \(l1 \ 0 \ 0 \ u1 \ 0 \ l2) \(l1 \ 0 \ 0 \ u1 \ u2 \ 0) \(0 \ l1 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ u2 \ 0) \(u1 \ 0 \ 0 \ l2) \(0 \ l1 \ u2 \ 0) \(0 \ l1 \ 0 \ l2)" using ivl1 ivl2 by (metis le_cases) lemma case_ivl_zero [consumes 2, case_names ZeroZero ZeroPos ZeroNeg PosZero NegZero NegNeg NegPos PosNeg PosPos]: fixes l1 u1 l2 u2 :: real shows "l1 \ u1 \ l2 \ u2 \ ((l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ 0 \ l2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ u2 \ 0) \ P) \ ((0 \ l1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ u2 \ 0) \ P) \ ((u1 \ 0 \ 0 \ l2) \ P) \ ((0 \ l1 \ u2 \ 0) \ P) \ ((0 \ l1 \ 0 \ l2) \ P) \ P" using ivl_zero_case[of l1 u1 l2 u2] by auto lemma case_inf2[case_names PosPos PosNeg PosNum NegPos NegNeg NegNum NumPos NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pu_inf[case_names PosAny AnyPos NegNeg NegNum NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ P w1 w2) \ (w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pl_inf[case_names NegAny AnyNeg PosPos PosNum NumPos NumNum]: shows "\w1 w2 P. (w1 = NEG_INF \ P w1 w2) \ (w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma word_trichotomy[case_names Less Equal Greater]: fixes w1 w2 :: word shows "(w1 P w1 w2) \ (w1 = w2 \ P w1 w2) \ (w2 P w1 w2) \ P w1 w2" using signed.linorder_cases by auto lemma case_times_inf [case_names PosPos NegPos PosNeg NegNeg PosLo PosHi PosZero NegLo NegHi NegZero LoPos HiPos ZeroPos LoNeg HiNeg ZeroNeg AllFinite]: fixes w1 w2 P assumes pp:"(w1 = POS_INF \ w2 = POS_INF \ P w1 w2)" and np:"(w1 = NEG_INF \ w2 = POS_INF \ P w1 w2)" and pn:"(w1 = POS_INF \ w2 = NEG_INF \ P w1 w2)" and nn:"(w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2)" and pl:"(w1 = POS_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and ph:"(w1 = POS_INF \ w2 \ POS_INF \ 0 P w1 w2)" and pz:"(w1 = POS_INF \ w2 = 0 \ P w1 w2)" and nl:"(w1 = NEG_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and nh:"(w1 = NEG_INF \ w2 \ POS_INF \ 0 P w1 w2)" and nz:"(w1 = NEG_INF \ 0 = w2 \ P w1 w2)" and lp:"(w1 \ NEG_INF \ w1 w2 = POS_INF \ P w1 w2)" and hp:"(w1 \ POS_INF \ 0 w2 = POS_INF \ P w1 w2)" and zp:"(0 = w1 \ w2 = POS_INF \ P w1 w2)" and ln:"(w1 \ NEG_INF \ w1 w2 = NEG_INF \ P w1 w2)" and hn:"(w1 \ POS_INF \ 0 w2 = NEG_INF \ P w1 w2)" and zn:"(0 = w1 \ w2 = NEG_INF \ P w1 w2)" and allFinite:"w1 \ NEG_INF \ w1 \ POS_INF \ w2 \ NEG_INF \ w2 \ POS_INF \ P w1 w2" shows " P w1 w2" proof (cases rule: word_trichotomy[of w1 0]) case Less then have w1l:"w1 Trivial arithmetic lemmas\ lemma max_diff_pos:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto lemma max_less:"2 ^ 31 < (9223372039002259455::int)" by auto lemma sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto lemma sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto lemma upcast_max:"sint((scast(0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto lemma upcast_min:"(0xFFFFFFFF80000001::64 Word.word) = ((scast (-0x7FFFFFFF::word))::64 Word.word)" by auto lemma min_extend_neg:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto lemma min_extend_val':"sint ((-0x7FFFFFFF)::64 Word.word) = (-0x7FFFFFFF)" by auto lemma min_extend_val:"(-0x7FFFFFFF::64 Word.word) = 0xFFFFFFFF80000001" by auto lemma range2s:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto section \Arithmetic operations\ text\This section defines operations which conservatively compute upper and lower bounds of arithmetic functions given upper and lower bounds on their arguments. Each function comes with a proof that it rounds in the advertised direction. \ subsection \Addition upper bound\ text\Upper bound of w1 + w2\ fun pu :: "word \ word \ word" where "pu w1 w2 = (if w1 = POS_INF then POS_INF else if w2 = POS_INF then POS_INF else if w1 = NEG_INF then (if w2 = NEG_INF then NEG_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum)) else if w2 = NEG_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" lemma scast_down_range: fixes w::"'a::len Word.word" assumes "sint w \ sints (len_of (TYPE('b::len)))" shows "sint w = sint ((scast w)::'b Word.word)" unfolding scast_def by (simp add: assms word_sint.Abs_inverse) lemma pu_lemma: fixes w1 w2 fixes r1 r2 :: real assumes up1:"w1 \\<^sub>U (r1::real)" assumes up2:"w2 \\<^sub>U (r2::real)" shows "pu w1 w2 \\<^sub>U (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word) = sint ((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2Geq:"sint ((scast w2)::64 Word.word) \ - (2 ^ 31) " using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 Word.word_size scast_eq1 upcast_max scast_eq3 len32 mem_Collect_eq by auto have "sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using Word.word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w2Less:"sint ((scast w2)::64 Word.word) < 9223372039002259455" by auto have w2Range: "-(2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 max_less) using max_diff_pos max_less w2Less w2Geq by auto have w2case1a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have w1Lower:"sint ((scast w1)::64 Word.word) \ - (2 ^ 31) " using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 Word.word_size scast_eq1 scast_eq2 scast_eq3 len32 mem_Collect_eq by auto have w1Leq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using Word.word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w1Less:"sint ((scast w1)::64 Word.word) < 9223372039002259455" using max_less by auto have w1MinusBound:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (-0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "0x80000001::32 Word.word"] Word.word_sint.Rep[of "(scast w1)::64 Word.word"] Word.word_sint.Rep[of "-0x7FFFFFFF::64 Word.word"] sints64 sints32) using w1Lower w1Less by auto have w1case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by (rule signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(- 0x7FFFFFFF)", OF w1MinusBound]) have w1case1a':"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val w1case1a by auto have w1Leq':"sint w1 \ 2^31 - 1" using Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using up1 up2 unfolding repU_def by auto show ?thesis proof (cases rule: case_pu_inf[where ?w1.0=w1, where ?w2.0=w2]) case PosAny then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case AnyPos then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case NegNeg then show ?thesis using up1 up2 by (auto simp add: repU_def repe.simps) next case NegNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = NEG_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast NEG_INF)::64 Word.word" have leq1:"r'\<^sub>1 \ (real_of_int (sint NEG_INF))" using equiv1 neq1 eq2 neq3 by (auto simp add: repe.simps) have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 by (auto simp add: repe.simps) have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" using up1 up2 apply (simp add: repU_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using w2case1a min_extend_neg apply (auto simp add: neq1 eq2 neq3 repINT repU_def repe.simps repeInt_simps up2 word_sless_alt) using repINT repU_def repe.simps repeInt_simps up2 word_sless_alt add.right_neutral add_mono dual_order.trans of_int_le_0_iff scast_eq3 by fastforce+ have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply(simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>2 - 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume " \ sint (scast w2 + 0xFFFFFFFF80000001) \ - 2147483647" have bound1:"r1 \ - 2147483647" using leq1 geq1 by (auto) have bound2:"r2 \ r'\<^sub>2" using leq2 geq2 by auto show "r1 + r2 \ r'\<^sub>2 - 2147483647" using bound1 bound2 by(linarith) qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - assume a:"\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have case1a:" sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" using case1a scast_eq3 min_extend_val' Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto have c:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val case1a by auto show "r'\<^sub>2 - 2147483647 = (real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word)))" using a b min_extend_val' scast_eq3 leq2 case1a by auto qed subgoal proof - have range2a:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 sints32 max_less) using max_diff_pos max_less w2Geq w2Less by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF range2a]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto assume "\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have b:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have d:"sint w2 \ 2^31 - 1" using Word.word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a b min_extend_val' using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have "sint (scast (((scast w2)::64 Word.word) + (-0x7FFFFFFF))::word) < 2147483647" unfolding downcast a b min_extend_val' using range2s[of "sint w2", OF d] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" by auto qed subgoal proof - assume notLeq:"\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have gr:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) from neg64 have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((-0x7FFFFFFF)::64 Word.word) = -0x7FFFFFFF" by auto have d:"sint w2 \ 2^31 - 1" using Word.word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) " using downcast by auto show "-2147483647 < real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word))" unfolding sintEq using gr of_int_less_iff of_int_minus of_int_numeral by linarith qed done have castEquiv:"\(?sum <=s scast NEG_INF) \ (scast ?sum) \\<^sub>U r1 + r2" using up1 up2 case1 case2 by fastforce have letRep:"(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" using case1 case2 by(cases "?sum <=s scast NEG_INF"; auto) show "pu w1 w2 \\<^sub>U r1 + r2" using letRep eq2 neq1 by(auto) next case NumNeg assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = NEG_INF" let ?sum = "(scast w1 + scast NEG_INF)::64 Word.word" have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" using up1 up2 apply (simp add: repU_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using w1case1a min_extend_neg apply (auto simp add: neq1 eq2 neq3 repINT repU_def repe.simps repeInt_simps up2 word_sless_alt) using repINT repU_def repe.simps repeInt_simps up2 word_sless_alt proof - fix r' assume a1:"sint ((scast w1)::64 Word.word) \ 0" then have h3:"sint w1 \ 0" using scast_eq1 by auto assume a2:"r2 \ r'" assume a3:"r1 \ (real_of_int (sint w1))" assume a4:"r' \ (- 2147483647)" from a2 a4 have h1:"r2 \ - 2147483647" by auto from a1 a3 h3 have h2:"r1 \ 0" using dual_order.trans of_int_le_0_iff by blast show "r1 + r2 \ (- 2147483647)" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint NEG_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w1MinusBound]) have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply (simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>1 - 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal using leq1 leq2 geq1 geq2 by auto apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding w1case1a using w2bound Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] scast_eq1) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto have "(real_of_int (sint ((scast w1)::64 Word.word ) - 2147483647)) = r'\<^sub>1 - (real_of_int 2147483647)" by (simp add: scast_eq1 leq2) then show "r'\<^sub>1 - 2147483647 = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word)))" by (metis b w1case1a' min_extend_val' diff_minus_eq_add minus_minus of_int_numeral) qed subgoal proof - assume "\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by(auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" using downcast min_extend_val' w1case1a' scast_eq1 arith[of "sint w1", OF w1Leq'] by auto qed subgoal proof - assume notLeq:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have gr:"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) show "- 2147483647 < real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word))" using scast_down_range[OF rightSize] gr of_int_less_iff of_int_minus of_int_numeral by auto qed done have letUp:"(let sum=?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1+r2" using case1 case2 by (auto simp add: Let_def) have puSimp:"pu w1 w2=(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum)" using neq3 neq1 eq2 equiv1 leq2 repeInt_simps by force then show "pu w1 w2 \\<^sub>U r1 + r2" using letUp puSimp by auto next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have inf_case:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>U r1 + r2" using repU_def repePos_simps by (meson dual_order.strict_trans not_less order_refl) have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "(w2)::32 Word.word"] Word.word_sint.Rep[of "(scast w1)::64 Word.word"] Word.word_sint.Rep[of "(scast w2)::64 Word.word"] sints64 sints32 by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ (- 0x7FFFFFFF)" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" then have sum_leq:"sint w1 + sint w2 \ - 0x7FFFFFFF" and sum_leq':" (sint w1 + sint w2) \ (- 2147483647)" using sint_eq unfolding Word.word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \scast w1 + scast w2 <=s - 0x7FFFFFFF\ word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (- 0x7FFFFFFF)" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans by auto qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" then have sum_leq:"sint w1 + sint w2 < 0x7FFFFFFF" unfolding Word.word_sle_def POS_INF_def using sint_eq by auto then have sum_leq':" (sint w1 + sint w2) < (2147483647)" by auto assume "\(?sum <=s ((scast NEG_INF)::64 Word.word))" then have sum_geq:"(- 0x7FFFFFFF) < sint w1 + sint w2" unfolding Word.word_sle_def NEG_INF_def using sint_eq by auto then have sum_geq':" (- 2147483647) < (sint w1 + sint w2)" by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) unfolding sint_eq using sints32 sum_geq sum_leq by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" using scast_down_range sints32 sum_geq sum_leq sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) show ?thesis using bound1 bound2 add_mono r12_sint_scast cast_eq sum_leq sum_leq' sum_geq' \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ by auto qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" proof (unfold repU_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (- 0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = real_of_int (sint w) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less (real_of_int (sint (uminus (minus(2 ^ 31) 1)))) (real_of_int (sint w))))" using leq anImp geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>U r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:"(-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repU_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have almost:"(let sum::64 Word.word = scast w1 + scast w2 in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ((?sum)::64 Word.word)") subgoal using inf_case Let_def int_case neg_inf_case by auto apply(cases "?sum <=s scast NEG_INF") subgoal using inf_case Let_def int_case neg_inf_case proof - assume "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" then have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ \ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using neg_inf_case by presburger then show ?thesis using int_case by force qed subgoal using inf_case Let_def int_case neg_inf_case proof - assume a1: "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" assume "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF" have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using a1 neg_inf_case by presburger then show ?thesis using int_case by force qed done then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by auto qed qed text\Lower bound of w1 + w2\ fun pl :: "word \ word \ word" where "pl w1 w2 = (if w1 = NEG_INF then NEG_INF else if w2 = NEG_INF then NEG_INF else if w1 = POS_INF then (if w2 = POS_INF then POS_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum)) else if w2 = POS_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" subsection \Addition lower bound\ text\Correctness of lower bound of w1 + w2\ lemma pl_lemma: assumes lo1:"w1 \\<^sub>L (r1::real)" assumes lo2:"w2 \\<^sub>L (r2::real)" shows "pl w1 w2 \\<^sub>L (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto have sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w2)) \ (-(2 ^ 31))" using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 mem_Collect_eq Word.word_size[of "(scast w2)::64 Word.word"] scast_eq1 scast_eq2 scast_eq3 len32 by auto then have thing4:"sint ((scast w2)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num using scast_eq3 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w2)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply ( auto simp add: Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using Word.word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto have thing3:" sint ((scast w2)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using lo1 lo2 unfolding repL_def by auto show ?thesis proof (cases rule: case_pl_inf[where ?w1.0=w1, where ?w2.0=w2]) case NegAny then show ?thesis apply (auto simp add: repL_def repe.simps) using lo1 lo2 linear by auto next case AnyNeg then show ?thesis apply (auto simp add: repL_def repe.simps) using linear by auto next case PosPos then show ?thesis using lo1 lo2 by (auto simp add: repL_def repe.simps) next case PosNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = POS_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast POS_INF)::64 Word.word" have case1:"(((scast POS_INF)::64 Word.word) <=s ?sum) \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) proof - fix r' assume a1:"0 \ sint (((scast w2)::64 Word.word))" from a1 have h3:"2147483647 \ sint w2 + 0x7FFFFFFF " using scast_eq3 by auto assume a2:"r' \ r1" assume a3:"(real_of_int (sint w2)) \ r2" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"2147483647 \ r1" by auto from a1 a3 h3 have h2:"0 \ r2" using dual_order.trans of_int_le_0_iff le_add_same_cancel2 by fastforce show "(2147483647) \ r1 + r2 " using h1 h2 h3 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>1 \ (real_of_int (sint POS_INF))" using equiv1 neq1 eq2 neq3 unfolding repe.simps by auto have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>2 + 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w2 + 0x7FFFFFFF)" have bound1:"2147483647 \ r1" using leq1 geq1 by (auto) have bound2:"r'\<^sub>2 \ r2 " using leq2 geq2 by auto show "r'\<^sub>2 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - assume a:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" by auto have case1a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have a1:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case1a by auto have c1:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "sint w2 + ( 0x7FFFFFFF) < 0x7FFFFFFF" using sintw2_bound case1a c1 scast_eq3 by linarith then have w2bound:"sint w2 < 0" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq3 c1 using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" by auto have c:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using case1a by auto have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have f:"r'\<^sub>2 = (real_of_int (sint w2))" by (simp add: leq2) show "r'\<^sub>2 + 2147483647 = (real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word)))" using a b c d scast_eq3 f leq2 of_int_numeral by fastforce qed subgoal proof - have truth2a:"-(2^(size ((scast w2)::64 Word.word)-1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2) using thing1 thing2 thing3 by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth2a]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto assume "\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" using neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have " 0x7FFFFFFF > sint w2 + ( 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:" sint w2 < 0" by simp have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 c apply (auto simp add: sints32 len32[of "TYPE(32)"]) using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] w2bound by auto have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" unfolding downcast a scast_eq3 c using w2bound by auto qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have gr:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) < 2147483647" by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" using neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "- 2147483647 \ w2" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w2" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w2" by auto have "- 2147483648 \ w2" apply(rule repe.cases[OF equiv2]) by auto then have "sint(- 2147483648::word) \ sint w2" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w2" by auto then have d:"sint w2 > - 2147483647" using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] neq3 n1 n2 by auto have w2bound:"- 2147483647 < sint w2 + 0x7FFFFFFF" using sintw2_bound case2a c scast_eq3 d by linarith have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" using sints32 len32[of "TYPE(32)"] w2bound Word.word_sint.Rep[of "(w2)::32 Word.word"] c case2a scast_eq3 sintw2_bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) " using downcast by auto show "- 2147483647 < real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word))" unfolding sintEq using gr of_int_less_iff of_int_minus of_int_numeral c case2a scast_eq3 w2bound by linarith qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using lo1 lo2 neq1 eq2 neq3 by (auto) next case NumPos assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = POS_INF" let ?sum = "(scast w1 + scast POS_INF)::64 Word.word" have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w1)) \ (-(2 ^ 31))" using Word.word_sint.Rep[of "(w1)::32 Word.word"] scast_eq1 scast_eq2 scast_eq3 Word.word_size[of "(scast w1)::64 Word.word"] sints32 len32 mem_Collect_eq by auto then have thing4:"sint ((scast w1)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num scast_eq3 scast_eq1 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w1)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using Word.word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by clarsimp have thing3:" sint ((scast w1)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have g:"(0x7FFFFFFF::64 Word.word) = 0x7FFFFFFF" by auto have c:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using g case1a by blast have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have e:"sint ((scast w1)::64 Word.word) = sint w1" using scast_eq1 by blast have d2:"sint w1 \ 2^31 - 1" using Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have case1:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) (* close 4 goals *) proof - fix r' have h3:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint (((scast w1)::64 Word.word)) + sint(0x7FFFFFFF::64 Word.word)" using case1a by auto have h4:"sint(0x7FFFFFFF::64 Word.word) = 2147483647" by auto assume a1:"0 \ sint ((scast w1)::64 Word.word)" then have h3:"sint w1 \ 0" using scast_eq1 h3 h4 by auto assume a2:"r' \ r2" assume a3:"(real_of_int (sint w1)) \ r1" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"r2 \ 2147483647" by auto from a1 a3 h3 have h2:"r1 \ 0" using dual_order.trans of_int_0_le_iff by fastforce show " 2147483647 \ r1 + r2" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint POS_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have neg64:"(((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w1)::64 Word.word) + (-0x7FFFFFFF)" by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>1 + 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w1 + 0x7FFFFFFF)" have bound1:"r2 \ 2147483647" using leq1 geq2 by (auto) have bound2:"r1 \ r'\<^sub>1" using leq2 geq1 by auto show "r'\<^sub>1 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w1)::64 Word.word) + (0x7FFFFFFF))" by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using w2bound Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] ) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint(((scast w1)::64 Word.word) + 0x7FFFFFFF)" using g by auto show "r'\<^sub>1 + 2147483647 = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word)))" using a b c d e f proof - have "(real_of_int (sint ((scast w1)::64 Word.word ) + 2147483647)) = r'\<^sub>1 + (real_of_int 2147483647)" using e leq2 by auto then show ?thesis using b c d of_int_numeral by metis qed qed subgoal proof - assume "\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e c using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w1)::64 Word.word) + ((0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" using downcast d e c arith[of "sint w1", OF d2] sintw2_bound by linarith qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have gr:"2147483647 > sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 using notLeq by auto have "0x7FFFFFFF > sint w1 + ( 0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have useful:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] useful by auto have "- 2147483647 \ w1" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w1" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w1" by auto have "- 2147483648 \ w1" apply(rule repe.cases[OF equiv1]) using int_not_undef[of w1] by auto then have "sint(- 2147483648::word) \ sint w1" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w1" by auto then have d:"sint w1 > - 2147483647" using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] n1 n2 neq3 by (simp) have d2:"sint (0x7FFFFFFF::64 Word.word) > 0" by auto from d d2 have d3:"- 2147483647 < sint w1 + sint (0x7FFFFFFF::64 Word.word)" by auto have d4:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint w1 + sint (0x7FFFFFFF::64 Word.word)" using case1a rightSize scast_down_range scast_eq1 by fastforce then show "-2147483647 < real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word))" using d3 d4 by auto qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using neq1 eq2 neq3 by (auto) next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 sints64 sints32 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "(w2)::32 Word.word"] by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ -0x7FFFFFFF" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" then have sum_leq:"sint w1 + sint w2 \ - 0x7FFFFFFF" and sum_leq':" (sint w1 + sint w2) \ (- 2147483647)" using sint_eq unfolding Word.word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using bound1 bound2 \scast w1 + scast w2 <=s -0x7FFFFFFF\ word_sle_def notinf1 notinf2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (-0x7FFFFFFF)" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans by auto qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" then have sum_leq:"sint w1 + sint w2 < 0x7FFFFFFF" unfolding Word.word_sle_def using sint_eq by auto then have sum_leq':" (sint w1 + sint w2) < (2147483647)" by auto assume "\(?sum <=s ((scast NEG_INF)::64 Word.word))" then have sum_geq:"(- 0x7FFFFFFF) < sint w1 + sint w2" unfolding Word.word_sle_def using sint_eq by auto then have sum_geq':" (- 2147483647) < (sint w1 + sint w2)" by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) unfolding sint_eq using sints32 sum_geq sum_leq by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" using scast_down_range sints32 sum_geq sum_leq sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) have leq_ref:"\x y ::real. x = y ==> x \ y" by auto show ?thesis apply(rule exI[where x="r'\<^sub>1 + r'\<^sub>2"]) apply(rule conjI) subgoal using r12_sint_scast cast_eq leq_ref r2w1 r1w1 add_mono[of r'\<^sub>1 r1 r'\<^sub>2 r2] bound1 bound2 by auto using bound1 bound2 add_mono r12_sint_scast cast_eq sum_leq sum_leq' sum_geq' sum_geq \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ by auto qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>L r1 + r2" proof (unfold repL_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (-0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ (real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word)))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = (real_of_int (sint w)) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less ( (real_of_int (sint (uminus (minus(2 ^ 31) 1))))) ((real_of_int (sint w)))))" using leq geq by auto qed have bigThree:"0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) \ \r'\r1 + r2. 2147483647 \ r'" proof - assume "0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have sum_leq:"0x7FFFFFFF \ sint w1 + sint w2 " and sum_leq':" 2147483647 \ (sint w1 + sint w2)" using sint_eq unfolding Word.word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \ 0x7FFFFFFF <=s scast w1 + scast w2 \ word_sle_def notinf1 notinf2 bound1 bound2 repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono of_int_add by fastforce show "\r'\ r1 + r2. (2147483647) \ r'" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans proof - have f1: " (real_of_int (sint w2)) = r'\<^sub>2" by (metis bound2 notinf2 notneginf2 repe.cases) have " (real_of_int (sint w1)) = r'\<^sub>1" by (metis bound1 notinf1 notneginf1 repe.cases) then have f2: " (real_of_int 2147483647) \ r'\<^sub>2 + r'\<^sub>1" using f1 sum_leq' by force have "r'\<^sub>2 + r'\<^sub>1 \ r2 + r1" by (meson add_left_mono add_right_mono bound1 bound2 order.trans) then show "r'\<^sub>1 + r'\<^sub>2 \ r1 + r2 \ 2147483647 \ r'\<^sub>1 + r'\<^sub>2" using f2 by (simp add: add.commute) qed qed have inf_case:"((scast POS_INF)::64 Word.word) <=s ?sum \ POS_INF \\<^sub>L r1 + r2" proof - assume "((scast POS_INF)::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have "((scast ((2 ^ 31 - 1)::word))::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" unfolding repL_def repe.simps by auto then have "(0x7FFFFFFF::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" by auto then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(0x7FFFFFFF \ r')" using bigThree by auto show "?thesis" unfolding repL_def repe.simps using leq geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>L r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:" (-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repL_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>L r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ?sum") apply(cases "?sum <=s scast NEG_INF") using inf_case neg_inf_case int_case by (auto simp add: Let_def) then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by(auto) qed qed subsection \Max function\ text\Maximum of w1 + w2 in 2s-complement\ fun wmax :: "word \ word \ word" where "wmax w1 w2 = (if w1 Correctness of wmax\ lemma wmax_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmax w1 w2 \\<^sub>E (max r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos from PosPos eq1 eq2 have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqInf:"wmax w1 w2 = POS_INF" using PosPos unfolding wmax.simps by auto have pos_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by linarith show ?thesis using pos_eq eqInf by auto next case PosNeg from PosNeg have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using eq1 eq2 by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding eq1 eq2 wmax.simps PosNeg word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum from PosNum eq1 eq2 have bound1:" (real_of_int (sint POS_INF)) \ r1" and bound2a:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" using PosNum bound2b unfolding wmax.simps word_sless_def word_sle_def by auto have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply (rule repPOS_INF) using bound1 bound2a bound2b word_sless_alt le_max_iff_disj unfolding eq1 eq2 by blast show "?thesis" using eqNeg neg_eq by auto next case NegPos from NegPos eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:" (real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding NegPos word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNeg from NegNeg eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"NEG_INF \\<^sub>E max r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 by(auto) have neg_eq:"wmax w1 w2 = NEG_INF" using NegNeg by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNum from NegNum eq1 eq2 have eq3:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"max r1 r2 = (real_of_int (sint w2))" using NegNum assms(2) bound2a eq3 repeInt_simps bound1 bound2a bound2b by (metis less_irrefl max.bounded_iff max_def not_less) then have extra_eq:"(wmax w1 w2) = w2" using assms(2) bound2a eq3 NegNum repeInt_simps by (simp add: word_sless_alt) have neg_eq:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using extra_eq eq3 bound2a bound2b by(auto) show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq extra_eq by auto next case NumPos from NumPos eq1 eq2 have p2:"w2 = POS_INF" and eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = r2" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have res1:"wmax w1 w2 = POS_INF" using NumPos p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (simp add: word_sless_alt) have res3:"POS_INF \\<^sub>E max r1 r2" using repPOS_INF NumPos bound2 le_max_iff_disj by blast show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res3 by auto next case NumNeg from NumNeg eq1 eq2 have n2:"w2 = NEG_INF" and rw1:"r1 = (real_of_int (sint w1))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using bound1b bound2 NumNeg less_trans wmax.simps of_int_less_iff word_sless_alt rw1 antisym_conv2 less_imp_le max_def by metis have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using bound1a bound1b bound2 NumNeg leD leI less_trans n2 wmax.simps by (auto simp add: word_sless_alt) show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto next case NumNum from NumNum eq1 eq2 have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b by (simp add: max_def word_sless_alt) have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \max r1 r2 = (real_of_int (sint (wmax w1 w2)))\ eq2 min_less_iff_disj)+ show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto qed lemma max_repU1: assumes "w1 \\<^sub>U x" assumes "w2 \\<^sub>U y" shows "wmax w1 w2 \\<^sub>U x " using wmax_lemma assms repU_def by (meson le_max_iff_disj) lemma max_repU2: assumes "w1 \\<^sub>U y" assumes "w2 \\<^sub>U x" shows "wmax w1 w2 \\<^sub>U x" using wmax_lemma assms repU_def by (meson le_max_iff_disj) text\Product of w1 * w2 with bounds checking\ fun wtimes :: "word \ word \ word" where "wtimes w1 w2 = (if w1 = POS_INF \ w2 = POS_INF then POS_INF else if w1 = NEG_INF \ w2 = POS_INF then NEG_INF else if w1 = POS_INF \ w2 = NEG_INF then NEG_INF else if w1 = NEG_INF \ w2 = NEG_INF then POS_INF else if w1 = POS_INF \ w2 0 0 = w2 then 0 else if w1 = NEG_INF \ w2 0 0 = w2 then 0 else if w1 w2 = POS_INF then NEG_INF else if 0 w2 = POS_INF then POS_INF else if 0 = w1 \ w2 = POS_INF then 0 else if w1 w2 = NEG_INF then POS_INF else if 0 w2 = NEG_INF then NEG_INF else if 0 = w1 \ w2 = NEG_INF then 0 else (let prod::64 Word.word = (scast w1) * (scast w2) in if prod <=s (scast NEG_INF) then NEG_INF else if (scast POS_INF) <=s prod then POS_INF else (scast prod)))" subsection \Multiplication upper bound\ text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_lower: fixes x y::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "- 4611686018427387904 \ x * y" proof - let ?thesis = "- 4611686018427387904 \ x * y" have is_neg:"- 4611686018427387904 < (0::int)" by auto have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * (2147483648) \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ y" using a6 by auto have h3:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * (-2147483648) \ 2147483648 * y" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ x" using a5 by auto have h3:"2147483648 * y \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg mult_right_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed show ?thesis using case1 case2 case3 case4 by linarith qed text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_upper: fixes x y ::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "x * y \ 4611686018427387904" proof - let ?thesis = "x * y \ 4611686018427387904" have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 by (simp add: mult_mono) show ?thesis using h1 h2 by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos2 by blast show ?thesis using h1 h2 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos by blast show ?thesis using h1 h2 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * -2147483648 \ x * -2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * -2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 mult_left_mono_neg by blast show ?thesis using h1 h2 by auto qed show "x * y \ 4611686018427387904" using case1 case2 case3 case4 by linarith qed text\Correctness of 32x32 bit multiplication\ subsection \Exact multiplication\ lemma wtimes_exact: assumes eq1:"w1 \\<^sub>E r1" assumes eq2:"w2 \\<^sub>E r2" shows "wtimes w1 w2 \\<^sub>E r1 * r2" proof - have POS_cast:"sint ((scast POS_INF)::64 Word.word) = sint POS_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have POS_sint:"sint POS_INF = (2^31)-1" by auto have w1_cast:"sint ((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2_cast:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have NEG_cast:"sint ((scast NEG_INF)::64 Word.word) = sint NEG_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have rangew1:"sint ((scast w1)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w1_cast by auto have rangew2:"sint ((scast w2)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w2_cast by auto show ?thesis proof (cases rule: case_times_inf[of w1 w2]) case PosPos then have a1: "PosInf \ r1" and a2: "PosInf \ r2" using "PosPos" eq1 eq2 repe.simps by (auto) have f3: "\n e::real. 1 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) have "\n e::real. 0 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) then have "r1 \ r1 * r2" using f3 "PosPos" eq1 eq2 repe.simps using eq1 eq2 by (auto simp add: repe.simps) then have "PosInf \ r1 * r2" using a1 by linarith then show ?thesis using "PosPos" by(auto simp add: repe.simps) next case NegPos from "NegPos" have notPos:"w1 \ POS_INF" unfolding POS_INF_def NEG_INF_def by auto have a1: "NegInf \ r1" using eq1 "NegPos" by (auto simp add: repe.simps) have a2: "PosInf \ r2" using eq2 "NegPos" by (auto simp add: repe.simps) have f1: "real_of_int Numeral1 = 1" by simp have f3: "(real_of_int 3) \ - r1" using a1 by (auto) have f4:"0 \ r2" using f1 a2 by(auto) have f5: "r1 \ - 1" using f3 by auto have fact:"r1 * r2 \ - r2" using f5 f4 mult_right_mono by fastforce show ?thesis using a1 a2 fact by (auto simp add: repe.simps "NegPos") next case PosNeg have a1: "PosInf \ r1" using eq1 "PosNeg" by (auto simp add: repe.simps) then have h1:"r1 \ 1" by (auto) have a2: " NegInf \ r2" using eq2 "PosNeg" by (auto simp add: repe.simps) have f1: "\ NegInf * (- 1) \ 1" by (auto) have f2: "\e ea::real. (e * (- 1) \ ea) = (ea * (- 1) \ e)" by force then have f3: "\ 1 * (- 1::real) \ NegInf" using f1 by blast have f4: "r1 * (- 1) \ NegInf" using f2 a1 by(auto) have f5: "\e ea eb. (if (ea::real) \ eb then e \ eb else e \ ea) = (e \ ea \ e \ eb)" by force have " 0 * (- 1::real) \ 1" by simp then have "r1 * (- 1) \ 0" using f5 f4 f3 f2 by meson then have f6: "0 \ r1" by fastforce have "1 * (- 1) \ (- 1::real)" using f2 by force then have fact:"r2 \ (- 1)" using f3 a2 by fastforce have rule:"\c. c > 0 \ r1 \ c \ r2 \ -1 \ r1 * r2 \ -c" apply auto by (metis (no_types, hide_lams) f5 mult_less_cancel_left_pos mult_minus1_right neg_le_iff_le not_less) have "r1 * r2 \ NegInf" using "PosNeg" f6 fact rule[of PosInf] a1 by(auto) then show ?thesis using "PosNeg" by (auto simp add: repe.simps) next case NegNeg have a1: "(-2147483647) \ r1" using eq1 "NegNeg" by (auto simp add: repe.simps) then have h1:"r1 \ -1" using max.bounded_iff max_def one_le_numeral by auto have a2: " (-2147483647) \ r2" using eq2 "NegNeg" by (auto simp add: repe.simps) have f1: "\e ea eb. \ (e::real) \ ea \ \ 0 \ eb \ eb * e \ eb * ea" using mult_left_mono by metis have f2: "- 1 = (- 1::real)" by force have f3: " 0 \ (1::real)" by simp have f4: "\e ea eb. (ea::real) \ e \ \ ea \ eb \ \ eb \ e" by (meson less_le_trans not_le) have f5: " 0 \ (2147483647::real)" by simp have f6: "- (- 2147483647) = (2147483647::real)" by force then have f7: "- ( (- 2147483647) * r1) = (2147483647 * r1)" by (metis mult_minus_left) have f8: "- ( (- 2147483647) * (- 1)) = 2147483647 * (- 1::real)" by simp have " 2147483647 = - 1 * (- 2147483647::real)" by simp then have f9: "r1 \ (- 1) \ 2147483647 \ r1 * (- 2147483647)" using f8 f7 f5 f2 f1 by linarith have f10: "- 2147483647 = (- 2147483647::real)" by fastforce have f11: "- (r2 * 1 * (r1 * (- 1))) = r1 * r2" by (simp add: mult.commute) have f12: "r1 * (- 1) = - (r1 * 1)" by simp have "r1 * 1 * ( (- 2147483647) * 1) = (- 2147483647) * r1" by (simp add: mult.commute) then have f13: "r1 * (- 1) * ( (- 2147483647) * 1) = 2147483647 * r1" using f12 f6 by (metis (no_types) mult_minus_left) have " 1 * r1 \ 1 * (- 2147483647)" using a1 by (auto simp add: a1) then have " 2147483647 \ r1 * (- 1)" by fastforce then have "0 \ r1 * (- 1)" using f5 f4 by (metis) then have "r1 \ (- 1) \ - (r1 * 2147483647) \ - (r2 * 1 * (r1 * (- 1)))" by (metis a2 f11 h1 mult_left_mono_neg minus_mult_right mult_minus1_right neg_0_le_iff_le) then have "r1 \ (- 1) \ r1 * (- 2147483647) \ r2 * r1" using f11 f10 by (metis mult_minus_left mult.commute) then have fact:" 2147483647 \ r2 * r1" using f9 f4 by blast show ?thesis using a1 a2 h1 fact by (auto simp add: repe.simps "NegNeg" mult.commute) next case PosLo from "PosLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosLo" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 < 0" using "PosLo" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ -1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosLo" by (auto simp add: repe.simps) have f4: "r1 * (- 1) \ (- 2147483647)" using upper by (auto) have f5: "r2 \ (- 1)" using lower2 rw2 by simp have "0 < r1" using upper by (auto) have "\r. r < - 2147483647 \ \ r < r1 * - 1" using f4 less_le_trans by blast then have "r1 * (real_of_int (sint w2)) \ (- 2147483647)" using f5 f4 upper lower2 rw2 mult_left_mono by (metis \0 < r1\ dual_order.order_iff_strict f5 mult_left_mono rw2) then have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower2 rw2 by (auto) then show ?thesis using "PosLo" by (auto simp add: repe.simps) next case PosHi from "PosHi" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosHi" have upper:"(real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 > 0" using "PosHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ 1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosHi" by (auto) have "0 \ r1" using upper by (auto) then have "r1 \ r1 * r2" using rw2 lower2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "PosInf \ r1 * r2" using upper lower2 rw2 apply (auto) using \0 \ r1\ mult_mono mult_numeral_1_right numeral_One of_int_1_le_iff zero_le_one by metis then show ?thesis using "PosHi" by (auto simp add: repe.simps) next case PosZero from "PosZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosZero" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 = 0" using "PosZero" by (auto simp add: word_sless_def word_sle_def) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosZero" by auto have "0 = r1 * r2" using "PosZero" rw2 by auto then show ?thesis using "PosZero" by (auto simp add: repe.simps) next case NegHi have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" using "NegHi" by (auto) from eq1 "NegHi" have upper:"(real_of_int (sint NEG_INF)) \ r1 " by (auto simp add: repe.simps) have low:"sint w2 > 0" using "NegHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w2)) > 0" by auto then have lower2:"(real_of_int (sint w2)) \ 1" using low by (simp add: int_le_real_less) from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegHi" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegHi" by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a5 a4 h1 h2 h3 h6 h5 a5 dual_order.trans leD mult.right_neutral by (metis dual_order.order_iff_strict mult_less_cancel_left2) qed have prereqs:"r1 \ - 1" "1 \ (real_of_int (sint w2))" " (- 2147483647::real) \ - 1 " "r1 \ (-2147483647)" using rw1 rw2 "NegHi" lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r1" " (real_of_int (sint w2))" " (- 2147483647)"] prereqs by auto then show ?thesis using "NegHi" by (auto simp add: repe.simps) next case NegLo from "NegLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "NegLo" have upper:"(real_of_int (sint NEG_INF)) \ r1" by (auto simp add: repe.simps) have low:"sint w2 < 0" using "NegLo" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w2)) < 0" by auto then have lower2:"(real_of_int (sint w2)) \ -1" using low by (simp add: int_le_real_less) from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegLo" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegLo" by (auto) have hom:"(- 2147483647) = -(2147483647::real)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult.right_neutral by (metis mult.commute neg_0_less_iff_less real_mult_le_cancel_iff1) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed have prereqs:"- 2147483647 < (0::real)" " r1 \ - 2147483647" " (real_of_int (sint w2)) \ - 1" using rw1 rw2 "NegLo" by (auto simp add: word_sless_def word_sle_def lower2) have "2147483647 \ r1 * r2" using upper lower1 lower2 rw1 rw2 prereqs mylem[of "-2147483647" "r1" "(real_of_int (sint w2))"] by (auto simp add: word_sless_def word_sle_def) then show ?thesis using "NegLo" by (auto simp add: repe.simps) next case NegZero from "NegZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq2 "NegZero" have "r2 = 0" using repe.simps "NegZero" by (auto) then show ?thesis using "NegZero" by (auto simp add: repe.simps) next case LoPos from "LoPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 < 0" using "LoPos" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w1 \ -1" by auto from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps "LoPos" by (auto simp add: repe.simps) have f4: "r2 * (- 1) \ (- 2147483647)" using upper by(auto) have f5: "r1 \ (- 1)" using lower2 rw1 by simp have "0 < r2" using upper by(auto) then have "r2 * r1 \ r2 * (- 1)" by (metis dual_order.order_iff_strict mult_right_mono f5 mult.commute) then have "r2 * r1 \ (- 2147483647)" by (meson f4 less_le_trans not_le) then have "(real_of_int (sint w1)) * r2 \ (- 2147483647)" using f5 f4 rw1 less_le_trans not_le mult.commute rw1 by (auto simp add: mult.commute) then have "r1 * r2 \ NegInf" using rw1 by (auto) then show ?thesis using "LoPos" by (auto simp: repe.simps) next case HiPos from "HiPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "HiPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 > 0" using "HiPos" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) then have lower2:"sint w1 \ 1" by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "HiPos" by (auto simp add: repe.simps) have "0 \ r2" using upper by(auto) then have "r2 \ r2 * r1" using lower2 rw2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "2147483647 \ r1 * r2" using upper lower2 rw2 by (auto simp add: word_sless_def word_sle_def order_trans) then show ?thesis using "HiPos" by (auto simp add: repe.simps) next case ZeroPos from "ZeroPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroPos" have upper:" (real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroPos" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroPos" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroPos" by (auto simp add: repe.simps) next case ZeroNeg from "ZeroNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroNeg" have upper:"(real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroNeg" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroNeg" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroNeg" by (auto simp add: repe.simps) next case LoNeg from "LoNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoNeg" have upper:" (real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have low:"sint w1 < 0" using "LoNeg" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w1)) < 0" by auto then have lower2:"(real_of_int (sint w1)) \ -1" using low by (simp add: int_le_real_less) from eq1 have rw1:"r2 \ (real_of_int (sint w2))" using "LoNeg" upper by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "LoNeg" by (auto simp add: upper repe.simps) have hom:"(- 2147483647::real) = -(2147483647)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult_left_mono mult.right_neutral mult.commute by (metis dual_order.order_iff_strict mult_minus_right mult_zero_right neg_le_iff_le) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed have prereqs:"- 2147483647 < (0::real)" " r2 \ - 2147483647" " (real_of_int (sint w1)) \ - 1" using rw1 rw2 "LoNeg" lower2 by (auto simp add: word_sless_def word_sle_def lower2) have "2147483647 \ r1 * r2" using upper lower1 lower2 rw1 rw2 mylem[of "-2147483647" "r2" "(real_of_int (sint w1))"] prereqs by (auto simp add:word_sless_def word_sle_def mult.commute) then show ?thesis using "LoNeg" by (auto simp add: repe.simps) next case HiNeg from HiNeg have w1NotPinf:"w1 \ POS_INF" and w1NotNinf:"w1 \ NEG_INF" by (auto) have upper:" (real_of_int (sint NEG_INF)) \ r2 " using HiNeg eq2 by (auto simp add: repe.simps ) have low:"sint w1 > 0" using HiNeg apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w1)) > 0" by auto then have lower2:"(real_of_int (sint w1)) \ 1" using low by (simp add: int_le_real_less) from eq2 have rw1:"r2 \ (real_of_int (sint w2))" using repe.simps HiNeg by (simp add: upper) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps HiNeg by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a4 h1 h2 h3 h6 h5 a5 dual_order.trans less_eq_real_def by (metis mult_less_cancel_left1 not_le) qed have prereqs:"r2 \ - 1" "1 \ (real_of_int (sint w1))" " (- 2147483647) \ - (1::real )" "r2 \ (- 2147483647)" using rw1 rw2 HiNeg lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ - 2147483647" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r2" "(real_of_int (sint w1))" " (- 2147483647)"] prereqs by (auto simp add: mult.commute) then show ?thesis using HiNeg by(auto simp add: repe.simps) next case AllFinite let ?prod = "(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" consider (ProdNeg) "?prod <=s ((scast NEG_INF)::64 Word.word)" | (ProdPos) "(((scast POS_INF)::64 Word.word) <=s ?prod)" | (ProdFin) "\(?prod <=s ((scast NEG_INF)::64 Word.word)) \ \((scast POS_INF)::64 Word.word) <=s ?prod" by (auto) then show ?thesis proof (cases) case ProdNeg have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume "?prod <=s ((scast NEG_INF)::64 Word.word)" then have sint_leq:"sint ?prod \ sint ((scast NEG_INF)::64 Word.word)" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by (auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by (auto simp add: repe.simps) show ?thesis using AllFinite ProdNeg w1_cast w2_cast rw1 rw2 sint_leq apply (auto simp add: repe.simps) by (metis (no_types, hide_lams) eq3 of_int_le_iff of_int_minus of_int_mult of_int_numeral) next case ProdPos have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume cast:"((scast POS_INF)::64 Word.word) <=s ?prod" then have sint_leq:"sint ((scast POS_INF)::64 Word.word) \ sint ?prod" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps AllFinite neqs by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps AllFinite neqs by auto have prodHi:"r1 * r2 \ PosInf" using w1_cast w2_cast rw1 rw2 sint_leq apply(auto) by (metis (no_types, hide_lams) eq3 of_int_le_iff of_int_mult of_int_numeral) have infs:"SCAST(32 \ 64) NEG_INF 64) POS_INF" by (auto) have casted:"SCAST(32 \ 64) POS_INF <=s SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2" using cast by auto have almostContra:"SCAST(32 \ 64) NEG_INF 64) w1 * SCAST(32 \ 64) w2" using infs cast signed.order.strict_trans2 by blast have contra:"\(SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2 <=s SCAST(32 \ 64) NEG_INF)" using eq3 almostContra by auto have wtimesCase:"wtimes w1 w2 = POS_INF" using neqs ProdPos almostContra wtimes.simps AllFinite ProdPos by (auto simp add: repe.simps Let_def) show ?thesis using prodHi apply(simp only: repe.simps) apply(rule disjI1) apply(rule exI[where x= "r1*r2"]) apply(rule conjI) apply(rule wtimesCase) using prodHi by auto next case ProdFin have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto from ProdFin have a1:"\(?prod <=s ((scast NEG_INF)::64 Word.word))" by auto then have sintGe:"sint (?prod) > sint (((scast NEG_INF)::64 Word.word))" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce from ProdFin have a2:"\((scast POS_INF)::64 Word.word) <=s ?prod" by auto then have sintLe:"sint (((scast POS_INF)::64 Word.word)) > sint (?prod)" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by(auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by(auto simp add: repe.simps) from rw1 rw2 have "r1 * r2 = (real_of_int ((sint w1) * (sint w2)))" by simp have rightSize:"sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) \ sints (len_of TYPE(32))" using sintLe sintGe sints32 by (simp) have downcast:"sint ((scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)))::word) = sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have res_eq:"r1 * r2 = real_of_int(sint((scast (((scast w1)::64 Word.word)*((scast w2)::64 Word.word)))::word))" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ by (auto) have res_up:"sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word) < sint POS_INF" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast w1 * scast w2) < sint (scast POS_INF)\ of_int_eq_iff res_eq by presburger have res_lo:"sint NEG_INF < sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word)" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast NEG_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast NEG_INF) < sint (scast w1 * scast w2)\ of_int_eq_iff res_eq by presburger have "scast ?prod \\<^sub>E (r1 * r2)" using res_eq res_up res_lo by (auto simp add: rep_simps) then show ?thesis using AllFinite ProdFin by(auto) qed qed qed subsection \Multiplication upper bound\ text\Upper bound of multiplication from upper and lower bounds\ fun tu :: "word \ word \ word \ word \ word" where "tu w1l w1u w2l w2u = wmax (wmax (wtimes w1l w2l) (wtimes w1u w2l)) (wmax (wtimes w1l w2u) (wtimes w1u w2u))" lemma tu_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E max (rl1 * rl2) (ru1 * rl2)" by (rule wmax_lemma[OF timesll timesul]) have maxt34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E max (rl1 * ru2) (ru1 * ru2)" by (rule wmax_lemma[OF timeslu timesuu]) have bigMax:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" by (rule wmax_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>U max (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repU_def by blast obtain maxt34val :: real where maxU34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>U max (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repU_def by blast obtain bigMaxU:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>U max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repU_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp add: mult_right_mono) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 repU_def timesuu tu.simps max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis wmax.elims) next case 2 have g1:"ru1 * ru2 \ 0" using "2" geq1 geq2 grl2 gru2 by (simp) have g2:"0 \ r1 * r2" using "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 3 have g1:"ru1 * ru2 \ 0" using "3" geq1 geq2 by simp have g2:"0 \ r1 * r2" using "3" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] repU_def tu.simps timesuu by (metis max.coboundedI1 max.commute maxt34) next case 4 have g1:"rl1 * rl2 \ rl1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 using \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ less_eq_real_def by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ by (metis mult_left_mono_neg mult.commute) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by (auto) then show ?thesis proof (cases) case 1 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by (simp add: mult_less_0_iff less_le_trans not_less) have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have case1:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_less_0_iff less_le_trans not_less by metis have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU2 max_repU1 repU_def timesll tu.simps) qed have case2:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute by (metis mult_left_mono_neg) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using case1 case2 le_cases by blast next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 have g1:"ru1 * ru2 \ 0" using r1 bounds grl2 gru2 gru1 leD leI less_le_trans by auto have g2:"0 \ r1 * r2" using r1 "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_less_0_iff not_less by metis have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesul tu.simps) next case 2 have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 less_eq(1) less_le_trans not_less mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 mult.commute not_le lower mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_eq(1) mult_le_cancel_left less_le_trans not_less by metis have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute not_le lower1 lower2 mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper1 lower2 mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono gru1 by metis from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis max_repU1 repU_def timesul tu.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper2 lower1 mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis repU_def tu.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF bigMaxU] max_repU2[OF maxU12] max_repU2[OF maxU34] by (metis repU_def tu.simps) qed qed subsection \Minimum function\ text\Minimum of 2s-complement words\ fun wmin :: "word \ word \ word" where "wmin w1 w2 = (if w1 Correctness of wmin\ lemma wmin_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmin w1 w2 \\<^sub>E (min r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos assume p1:"w1 = POS_INF" and p2:"w2 = POS_INF" then have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" using eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqInf:"wmin w1 w2 = POS_INF" using p1 p2 unfolding wmin.simps by auto have pos_eq:"POS_INF \\<^sub>E min r1 r2" apply(rule repPOS_INF) using bound1 bound2 unfolding eq1 eq2 by auto show ?thesis using pos_eq eqInf by auto next case PosNeg assume p1:"w1 = POS_INF" assume n2:"w2 = NEG_INF" obtain r ra :: real where bound1:" (real_of_int (sint POS_INF)) \ r" and bound2:"ra \ (real_of_int (sint NEG_INF))" and eq1:"r1 = r" and eq2:"r2 = ra" using p1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps p1 n2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum assume p1:"w1 = POS_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2a:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" using p1 np2 nn2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"min r1 r2 = sint w2" using p1 by (metis bound1 bound2b dual_order.trans eq2 min_def not_less) have neg_eq:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1 bound2a bound2b bound2b p1 unfolding eq1 eq2 by (auto simp add: word_sless_alt) show "?thesis" using eqNeg neg_eq by (metis bound2b less_eq_real_def not_less of_int_less_iff p1 wmin.simps word_sless_alt) next case NegPos assume n1:"w1 = NEG_INF" assume p2:"w2 = POS_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"(real_of_int (sint POS_INF)) \ r2" using n1 p2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps n1 p2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 unfolding eq1 eq2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNeg assume n1:"w1 = NEG_INF" assume n2:"w2 = NEG_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using n1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 unfolding NEG_INF_def by (auto) have neg_eq:"wmin w1 w2 = NEG_INF" using n1 n2 unfolding NEG_INF_def wmin.simps by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNum assume n1:"w1 = NEG_INF" and nn2:"w2 \ NEG_INF" and np2:"w2 \ POS_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" using n1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" using n1 assms(2) bound2a eq2 n1 repeInt_simps by (auto simp add: word_sless_alt) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2a bound2b eq1 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NumPos assume p2:"w2 = POS_INF" and nn1:"w1 \ NEG_INF" and np1:"w1 \ POS_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:" (real_of_int (sint POS_INF)) \ r2" using nn1 np1 p2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = w1" using p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (auto simp add: word_sless_alt) have res2:"min r1 r2 = (real_of_int (sint w1))" using eq1 eq2 bound1a bound1b bound2 by (auto simp add: less_imp_le less_le_trans min_def) have res3:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply(rule repINT) using p2 bound1a res1 bound1a bound1b bound2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 res3 by auto next case NumNeg assume nn1:"w1 \ NEG_INF" assume np1:"w1 \ POS_INF" assume n2:"w2 = NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using nn1 np1 n2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = NEG_INF" using n2 bound1b by (metis min.absorb_iff2 min_def n2 not_less of_int_less_iff wmin.simps word_sless_alt) have res2:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1a bound1b bound2 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto next case NumNum assume np1:"w1 \ POS_INF" assume nn1:"w1 \ NEG_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" using nn1 np1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"min r1 r2 = (real_of_int (sint (wmin w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b by (simp add: min_def word_sless_alt) have res2:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \min r1 r2 = (real_of_int (sint (wmin w1 w2)))\ eq2 min_less_iff_disj)+ show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto qed lemma min_repU1: assumes "w1 \\<^sub>L x" assumes "w2 \\<^sub>L y" shows "wmin w1 w2 \\<^sub>L x " using wmin_lemma assms repL_def by (meson min_le_iff_disj) lemma min_repU2: assumes "w1 \\<^sub>L y" assumes "w2 \\<^sub>L x" shows "wmin w1 w2 \\<^sub>L x" using wmin_lemma assms repL_def by (meson min_le_iff_disj) subsection \Multiplication lower bound\ text\Multiplication lower bound\ fun tl :: "word \ word \ word \ word \ word" where "tl w1l w1u w2l w2u = wmin (wmin (wtimes w1l w2l) (wtimes w1u w2l)) (wmin (wtimes w1l w2u) (wtimes w1u w2u))" text\Correctness of multiplication lower bound\ lemma tl_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E min (rl1 * rl2) (ru1 * rl2)" by (rule wmin_lemma[OF timesll timesul]) have maxt34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E min (rl1 * ru2) (ru1 * ru2)" by (rule wmin_lemma[OF timeslu timesuu]) have bigMax:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E min (min(rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" by (rule wmin_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>L min (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repL_def by blast obtain maxt34val :: real where maxU34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>L min (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repL_def by blast obtain bigMaxU:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>L min (min (rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repL_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" and geq3:"rl1 \ 0" and geq4:"rl2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * ru2 \ 0" using "1" geq1 geq2 geq3 geq4 grl2 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 min_repU1 min_repU2 repL_def repU_def timeslu tl.simps wmin.elims by (metis bigMax min_le_iff_disj) next case 2 have g1:"rl1 * ru2 \ rl1 * r2" using "2" geq1 geq2 grl2 gru2 by (metis mult_le_cancel_left geq3 leD) have g2:"rl1 * r2 \ r1 * r2" using "2" geq1 geq2 grl2 gru2 by (simp add: mult_right_mono grl1) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 min_repU2 repL_def tl.simps min.coboundedI1 maxt34) next case 3 have g1:"ru1 * rl2 \ ru1 * r2" using "3" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "3" geq1 geq2 grl1 grl2 gru1 gru2 mult_minus_right mult_right_mono by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 timesul by (metis repL_def tl.simps) next case 4 have g1:"ru1 * rl2 \ 0" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ mult_less_0_iff less_eq_real_def not_less by auto have g2:"0 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 by (metis mult_less_0_iff not_less) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU1 min_repU2 repL_def timesul tl.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) have g2:"0 \ r1 * r2" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis by (metis repL_def timesll tl.simps up maxU12 maxU34 wmin.elims min_repU2 min_repU1) next case 2 have bound:"ru2 \ 0" using "2" r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"rl1 * ru2 \ rl1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have g2:"rl1 * r2 \ r1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff mult_le_cancel_right by fastforce from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 wmin.elims min_repU2 min.coboundedI1 maxt34 repL_def tl.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast consider (Pos) "r1 \ 0" | (Neg) "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case Pos have bound:"rl2 \ 0" using Pos r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"ru1 * rl2 \ ru1 * r2" using Pos bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have p1:"\a::real. (0 \ - a) = (a \ 0)" by(auto) have p2:"\a b::real. (- a \ - b) = (b \ a)" by auto have g2:"ru1 * r2 \ r1 * r2" using Pos r2 bounds grl1 grl2 gru1 gru2 p1 p2 by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU2 min_repU1 repL_def timesul tl.simps) next case Neg have g1:"ru1 * ru2 \ 0" using Neg r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using Neg r2 zero_le_mult_iff by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have bound:"0 \ ru1" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 zero_le_mult_iff by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt12 maxt34 repL_def timesll tl.simps by metis next case 2 have g1:"ru1 * rl2 \ ru1 * r2" using r1 "2" bounds bound grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 "2" bounds bound grl2 gru2 by (metis mult_left_mono_neg gru1 mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 by (metis repL_def timesul tl.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have bound:"rl1 \ 0" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 assume r2:"r2 \ 0" have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds bound grl1 grl2 gru1 gru2 by (metis mult_le_cancel_left leD) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt34 by (metis min_repU2 repL_def tl.simps) next case 2 assume r2:"r2 \ 0" have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 r2 by (simp add: zero_le_mult_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using not_less mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left mult.commute not_le lower1 lower2 by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_le_trans upper1 lower2 by (metis mult_le_cancel_left not_less) have g2:"rl1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono mult_le_0_iff grl1 by blast from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt12 maxt34 by (metis repL_def timeslu tl.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono less_le_trans not_less by metis have g2:"ru1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 by (metis repL_def timesul tl.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by auto have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt12 maxt34 by (metis repL_def tl.simps) qed qed text\Most significant bit only changes under successor when all other bits are 1\ lemma msb_succ: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0xFFFFFFFF" assumes neq2:"uint w \ 0x7FFFFFFF" shows "msb (w + 1) = msb w" proof - have "w \ 0xFFFFFFFF" using neq1 by auto then have neqneg1:"w \ -1" by auto have "w \ 0x7FFFFFFF" using neq2 by auto then have neqneg2:"w \ (2^31)-1" by auto show ?thesis using neq1 neq2 unfolding msb_big using Word_Lemmas.word_le_make_less[of "w + 1" "0x80000000"] Word_Lemmas.word_le_make_less[of "w " "0x80000000"] neqneg1 neqneg2 by auto qed text\Negation commutes with msb except at edge cases\ lemma msb_non_min: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0" assumes neq2:"uint w \ ((2^(len_of (TYPE(31)))))" shows "msb (uminus w) = HOL.Not(msb(w))" proof - have fact1:"uminus w = word_succ (~~ w)" by (rule twos_complement) have fact2:"msb (~~w) = HOL.Not(msb w)" using Word.word_ops_msb[of w] by auto have neqneg1:"w \ 0" using neq1 by auto have not_undef:"w \ 0x80000000" using neq2 by auto then have neqneg2:"w \ (2^31)" by auto have "~~ 0xFFFFFFFF = (0::word)" by auto then have "(~~ w) \ 0xFFFFFFFF" using neqneg1 neqneg2 Word.word_bool_alg.double_compl[of w] by metis then have uintNeq1:"uint (~~ w) \ 0xFFFFFFFF" using uint_distinct[of "~~w" "0xFFFFFFFF"] by auto - have "~~ (0x7FFFFFFF::word) = (2 ^ 31::word)" by (auto simp add: max_word_def) + have "~~ (0x7FFFFFFF::word) = (2 ^ 31::word)" + by simp then have "(~~ w) \ 0x7FFFFFFF" using neqneg1 neqneg2 Word.word_bool_alg.double_compl[of w] by metis then have uintNeq2:" uint (~~ w) \ 0x7FFFFFFF" using uint_distinct[of "~~w" "0x7FFFFFFF"] by auto have fact3:"msb ((~~w) + 1) = msb (~~w)" apply(rule msb_succ[of "~~w"]) using neq1 neq2 uintNeq1 uintNeq2 by auto show "msb (uminus w) = HOL.Not(msb(w))" using fact1 fact2 fact3 by (simp add: word_succ_p1) qed text\Only 0x80000000 preserves msb=1 under negation\ lemma msb_min_neg: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"msb w" shows "uint w = ((2^(len_of (TYPE(31)))))" proof - have neq1:"uint w \ 0" using msb2 word_msb_0 uint_0_iff by force show ?thesis using msb1 msb2 msb_non_min[OF neq1] by auto qed text\Only 0x00000000 preserves msb=0 under negation\ lemma msb_zero: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"\ msb w" shows "uint w = 0" proof - have neq:"w \ ((2 ^ len_of TYPE(31))::word)" using msb1 msb2 by auto have eq:"uint ((2 ^ len_of TYPE(31))::word) = 2 ^ len_of TYPE(31)" by auto then have neq:"uint w \ uint ((2 ^ len_of TYPE(31))::word)" using uint_distinct[of w "2^len_of TYPE(31)"] neq eq by auto show ?thesis using msb1 msb2 minus_zero msb_non_min[of w] neq by force qed text\Finite numbers alternate msb under negation\ lemma msb_pos: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"\ msb w" shows "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" proof - have main: "w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint apply(rule conjI) apply (metis neg_equal_0_iff_equal not_le word_less_1) proof - have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "msb w" using Word_Lemmas.msb_big[of w] by auto then show False using msb2 by auto qed have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" subgoal for w1 w2 by (simp add: word_le_def) done have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" subgoal for w1 w2 by (simp add: word_less_def) done have gr_to_geq:"w > 0x7FFFFFFF \ w \ 0x80000000" apply(rule mylem) using mylem2[of "0x7FFFFFFF" "w"] by auto have taut:"w \ 0x7FFFFFFF \ w > 0x7FFFFFFF" by auto then show "w \ 0x7FFFFFFF" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({1..(minus(2 ^ (minus(len_of TYPE(32)) 1)) 1)})::word set)) = ({1..minus(2 ^ (minus (len_of TYPE(32)) 1)) 1}::int set)" apply(auto simp add: word_le_def) subgoal for xa proof - assume lower:"1 \ xa" and upper:"xa \ 2147483647" then have in_range:"xa \ {0 .. 2^32-1}" by auto then have "xa \ range (uint::word \ int)" unfolding Word.word_uint.Rep_range Word.uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have "w \ {1..0x7FFFFFFF} " using lower upper apply(clarsimp, auto) by (auto simp add: word_le_def) then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {1..2 ^ (len_of TYPE(32) - 1) - 1}" using uint_distinct uint_distinct main image_eqI by blast qed lemma msb_neg: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"msb w" shows "uint w \ {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" proof - have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" by (simp add: word_le_def) have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" by (simp add: word_less_def) have gr_to_geq:"w > 0x80000000 \ w \ 0x80000001" apply(rule mylem) using mylem2[of "0x80000000" "w"] by auto have taut:"w \ 0x80000000 \ 0x80000000 < w" by auto have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "(msb (-w))" using Word_Lemmas.msb_big[of "-w"] Word_Lemmas.msb_big[of "w"] by (simp add: msb2) then show False using msb1 by auto qed have main: "w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint proof - show "0x80000001 \ w" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1})::word set)) = {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" apply(auto) subgoal for xa by (simp add: word_le_def) subgoal for w using Word.word.uint[of w] by auto subgoal for xa proof - assume lower:"2147483649 \ xa" and upper:"xa \ 4294967295" then have in_range:"xa \ {0x80000000 .. 0xFFFFFFFF}" by auto then have "xa \ range (uint::word \ int)" unfolding Word.word_uint.Rep_range Word.uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have the_in:"w \ {0x80000001 .. 0xFFFFFFFF} " using lower upper by (auto simp add: word_le_def) have the_eq:"(0xFFFFFFFF::word) = -1" by auto from the_in the_eq have "w \ {0x80000001 .. -1}" by auto then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using uint_distinct uint_distinct main image_eqI by blast qed text\2s-complement commutes with negation except edge cases\ lemma sint_neg_hom: fixes w :: "32 Word.word" shows "uint w \ ((2^(len_of (TYPE(31))))) \ (sint(-w) = -(sint w))" unfolding Word_Bitwise.word_sint_msb_eq apply auto subgoal using msb_min_neg by auto prefer 3 subgoal using msb_zero[of w] by (simp add: msb_zero) proof - assume msb1:"msb (- w)" assume msb2:"\ msb w" have "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb_pos[OF msb1 msb2] by auto then have bound:"uint w \ {1 .. 0x7FFFFFFF}" by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ ((- x) mod (2^n)) - (2^n) = - x" subgoal for x n apply(cases "x mod 2^n = 0") by(auto simp add: Divides.zmod_zminus1_eq_if[of x "2^n"]) done have lem_rule:"uint w \ {1..2 ^ 32 - 1} \ (- uint w mod 4294967296) - 4294967296 = - uint w" using lem[of "uint w" 32] by auto have almost:"- uint w mod 4294967296 - 4294967296 = - uint w" apply(rule lem_rule) using bound by auto show "uint (- w) - 2 ^ size (- w) = - uint w" using bound unfolding Word.uint_word_ariths word_size_neg by (auto simp add: size almost) next assume neq:"uint w \ 0x80000000" assume msb1:"\ msb (- w)" assume msb2:"msb w" have bound:"uint w \ {0x80000001.. 0xFFFFFFFF}" using msb1 msb2 msb_neg by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ (-x mod (2^n)) = (2^n) - x" subgoal for x n apply(auto) apply(cases "x mod 2^n = 0") by (simp add: Divides.zmod_zminus1_eq_if[of x "2^n"])+ done from bound have wLeq: "uint w \ 4294967295" and wGeq: "2147483649 \ uint w" by auto from wLeq have wLeq':"uint w \ 4294967296" by fastforce have f3: "(0 \ 4294967296 + - 1 * uint w + - 1 * ((4294967296 + - 1 * uint w) mod 4294967296)) = (uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296)" by auto have f4: "(0 \ 4294967296 + - 1 * uint w) = (uint w \ 4294967296)" by auto have f5: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * (i mod ia)" by (simp add: zmod_le_nonneg_dividend) then have f6: "uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296" using f4 f3 wLeq' by blast have f7: "4294967296 + - 1 * uint w + - 4294967296 = - 1 * uint w" by auto have f8: "- (1::int) * 4294967296 = - 4294967296" by auto have f9: "(0 \ - 1 * uint w) = (uint w \ 0)" by auto have f10: "(4294967296 + -1 * uint w + -1 * ((4294967296 + -1 * uint w) mod 4294967296) \ 0) = (4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296)" by auto have f11: "\ 4294967296 \ (0::int)" by auto have f12: "\x0. ((0::int) < x0) = (\ x0 \ 0)" by auto have f13: "\x0 x1. ((x1::int) < x0) = (\ 0 \ x1 + - 1 * x0)" by auto have f14: "\x0 x1. ((x1::int) \ x1 mod x0) = (x1 + - 1 * (x1 mod x0) \ 0)" by auto have "\ uint w \ 0" using wGeq by fastforce then have "4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296" using f14 f13 f12 f11 f10 f9 f8 f7 by (metis (no_types) int_mod_ge) then show "uint (- w) = 2 ^ size w - uint w" using f6 unfolding Word.uint_word_ariths by (auto simp add: size f4) qed text\2s-complement encoding is injective\ lemma sint_dist: fixes x y ::word assumes "x \ y" shows "sint x \ sint y" by (simp add: assms) subsection\Negation\ fun wneg :: "word \ word" where "wneg w = (if w = NEG_INF then POS_INF else if w = POS_INF then NEG_INF else -w)" text\word negation is correct\ lemma wneg_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wneg w \\<^sub>E -r" apply(rule repe.cases[OF eq]) apply(auto intro!: repNEG_INF repPOS_INF simp add: repe.simps)[2] subgoal for ra proof - assume eq:"w = ra" assume i:"r = (real_of_int (sint ra))" assume bounda:" (real_of_int (sint ra)) < (real_of_int (sint POS_INF))" assume boundb:" (real_of_int (sint NEG_INF)) < (real_of_int (sint ra))" have raNeq:"ra \ 2147483647" using sint_range[OF bounda boundb] by (auto) have raNeqUndef:"ra \ 2147483648" using int_not_undef[OF bounda boundb] by (auto) have "uint ra \ uint ((2 ^ len_of TYPE(31))::word)" apply (rule uint_distinct) using raNeqUndef by auto then have raNeqUndefUint:"uint ra \ ((2 ^ len_of TYPE(31)))" by auto have res1:"wneg w \\<^sub>E (real_of_int (sint (wneg w)))" apply (rule repINT) using sint_range[OF bounda boundb] sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndefUint raNeqUndef eq by(auto) have res2:"- r = (real_of_int (sint (wneg w)))" using eq bounda boundb i sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndef eq by (auto) show ?thesis using res1 res2 by auto qed done subsection\Comparison\ fun wgreater :: "word \ word \ bool" where "wgreater w1 w2 = (sint w1 > sint w2)" lemma neg_less_contra:"\x. Suc x < - (Suc x) \ False" by auto text\Comparison < is correct\ lemma wgreater_lemma:"w1 \\<^sub>L (r1::real) \ w2 \\<^sub>U r2 \ wgreater w1 w2 \ r1 > r2" proof (auto simp add: repU_def repL_def) fix r'\<^sub>1 r'\<^sub>2 assume sint_le:"sint w1 > sint w2" then have sless:"(w2 1 \ r1" assume r2_leq:"r2 \ r'\<^sub>2" assume wr1:"w1 \\<^sub>E r'\<^sub>1" assume wr2:"w2 \\<^sub>E r'\<^sub>2" have greater:"r'\<^sub>1 > r'\<^sub>2" using wr1 wr2 apply(auto simp add: repe.simps) prefer 4 using sless sint_le by (auto simp add: less_le_trans not_le) show "r1 > r2" using r1_leq r2_leq greater by auto qed text\Comparison $\geq$ of words\ fun wgeq :: "word \ word \ bool" where "wgeq w1 w2 = ((\ ((w2 = NEG_INF \ w1 = NEG_INF) \(w2 = POS_INF \ w1 = POS_INF))) \ (sint w2 \ sint w1))" text\Comparison $\geq$ of words is correct\ lemma wgeq_lemma:"w1 \\<^sub>L r1 \ w2 \\<^sub>U (r2::real) \ wgeq w1 w2 \ r1 \ r2" proof (unfold wgeq.simps) assume assms:"\ (w2 = NEG_INF \ w1 = NEG_INF \ w2 = POS_INF \ w1 = POS_INF) \ sint w2 \ sint w1" assume a1:"w1 \\<^sub>L r1" and a2:"w2 \\<^sub>U (r2::real)" from assms have sint_le:"sint w2 \ sint w1" by auto then have sless:"(w2 <=s w1)" using word_sless_alt word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 where r1_leq:"r'\<^sub>1 \ r1" and r2_leq:"r2 \ r'\<^sub>2" and wr1:"w1 \\<^sub>E r'\<^sub>1" and wr2:"w2 \\<^sub>E r'\<^sub>2" using a1 a2 unfolding repU_def repL_def by auto from assms have check1:"\ (w1 = NEG_INF \ w2 = NEG_INF)" by auto from assms have check2:"\ (w1 = POS_INF \ w2 = POS_INF)" by auto have less:"r'\<^sub>2 \ r'\<^sub>1" using sless sint_le check1 check2 repe.simps wr2 wr1 by(auto simp add: repe.simps) show "r1 \ r2" using r1_leq r2_leq less by auto qed subsection\Absolute value\ text\Absolute value of word\ fun wabs :: "word \ word" where "wabs l1 = (wmax l1 (wneg l1))" text\Correctness of wmax\ lemma wabs_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wabs w \\<^sub>E (abs r)" proof - have w:"wmax w (wneg w) \\<^sub>E max r (-r)" by (rule wmax_lemma[OF eq wneg_lemma[OF eq]]) have r:"max r (-r) = abs r" by auto from w r show ?thesis by auto qed end \ No newline at end of file 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,236 +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 "(max_word::16 word) = 65535" by(simp add: max_word_def) - hence port_range: "\s e port. s = 0 \ e = 0xFFFF \ (port::16 word) \ 0xFFFF" by simp + have "65535 = (max_word::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,501 +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 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) \ ProtoAny \ a \ a = {p. p \ ProtoAny}" proof - have all_proto_hlp: "ProtoAny \ a \ (\p \ {0..max_word}. 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)" - by(auto simp add: max_word_def word_upto_set_eq) + proof - + have \0xFF = (max_word :: 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)) \ ((\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 "(max_word::16 word) = 65535" by(simp add: max_word_def) - hence port_range: "\s e port. s = 0 \ e = 0xFFFF \ (port::16 word) \ 0xFFFF" by simp + have "65535 = (max_word::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(simp_all) - apply(simp_all add: max_word_def) + apply (simp_all add: simple_match_any_def match_ifaceAny ipset_from_cidr_0) + apply (subgoal_tac [!] "(65535::16 word) = max_word") + 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,1052 +1,1052 @@ 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)\" 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); 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" 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" 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 ?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) 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 ffff: "0xffff = word_of_int (2 ^ 16 - 1)" by fastforce - show ?thesis using max_word_max[of a] unfolding max_word_def ffff by fastforce + 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) (* 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 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 Some_to_the[symmetric];fail) 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/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) - apply(simp add: simple_match_any_def match_ifaceAny) - apply(subgoal_tac "(65535::16 word) = max_word") - apply(simp) - by(simp add: max_word_def) + apply (simp add: simple_conn_matches_def simple_match_any_def match_ifaceAny) + apply (subgoal_tac "(65535::16 word) = max_word") + 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 (meson build_ip_partition_same_fw_min int_not_emptyD same_fw_behaviour_one_equi(1)) 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,421 +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" by(simp add: max_word_def) - thus ?thesis by(simp add: simple_match_any_def ipset_from_cidr_0 match_ifaceAny) + have *: "(65535::16 word) = max_word" + 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(cases ifce) (simp add: match_iface_refl) } 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)" 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 ( 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)" 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/Word_Enum.thy b/thys/Word_Lib/Word_Enum.thy --- a/thys/Word_Lib/Word_Enum.thy +++ b/thys/Word_Lib/Word_Enum.thy @@ -1,95 +1,95 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Thomas Sewell *) section "Enumeration Instances for Words" theory Word_Enum imports Enumeration Word_Lib begin instantiation word :: (len) enum begin definition "(enum_class.enum :: ('a :: len) word list) \ map of_nat [0 ..< 2 ^ LENGTH('a)]" definition "enum_class.enum_all (P :: ('a :: len) word \ bool) \ Ball UNIV P" definition "enum_class.enum_ex (P :: ('a :: len) word \ bool) \ Bex UNIV P" instance apply (intro_classes) apply (force simp: enum_word_def) apply (simp add: distinct_map enum_word_def) apply (rule subset_inj_on, rule word_unat.Abs_inj_on) apply (clarsimp simp add: unats_def) apply (simp add: enum_all_word_def) apply (simp add: enum_ex_word_def) done end 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 length_word_enum: "length (enum :: 'a :: len word list) = 2 ^ LENGTH('a)" by (simp add: enum_word_def) 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) declare of_nat_diff [simp] 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) 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" - by (simp add: maxBound_word max_word_minus [symmetric]) + by (fact maxBound_word) lemma leq_maxBound [simp]: "(x::'a::len word) \ maxBound" by (simp add: maxBound_max_word) 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,6252 +1,6250 @@ (* * 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 Word_EqI Word_Enum "HOL-Library.Sublist" begin lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith 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 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 (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI) 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" by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size) lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp 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 same_length_is_parallel: assumes len: "\y \ set as. length y = x" shows "\x \ set as. \y \ set as - {x}. x \ y" proof (rule, rule) fix x y assume xi: "x \ set as" and yi: "y \ set as - {x}" from len obtain q where len': "\y \ set as. length y = q" .. show "x \ y" proof (rule not_equal_is_parallel) from xi yi show "x \ y" by auto from xi yi len' show "length x = length y" by (auto dest: bspec) qed qed text \Lemmas about words\ lemmas and_bang = word_and_nth lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x && mask (size x - n))" by (simp add: of_bl_drop word_size_bl) 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 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_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_unat.Rep_eqD) 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 set_enum_word8_def: "(set enum::word8 set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by eval lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: word8 shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done 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, unat_arith) apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms) apply (rule map_cong[OF _ refl]) apply (rule arg_cong2[where f = "\x y. [x ..< y]"], unat_arith) apply (rule refl) 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 (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply simp apply clarsimp apply (erule disjE) apply simp apply clarsimp apply (rule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply assumption apply simp apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]]) apply clarsimp apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply (rule conjI) apply (subst word_le_nat_alt [symmetric]) apply simp apply safe apply (simp add: word_le_nat_alt [symmetric]) apply (simp add: word_less_nat_alt [symmetric]) 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 by simp 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 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 mask_shift: "(x && ~~ (mask y)) >> y = x >> y" by word_eqI 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 shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_def) 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_unat.Rep_eqD) 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 (case_tac "size w \ n", clarsimp simp: shiftl_zero_size) (clarsimp simp: not_le ucast_bl bl_shiftl bang_eq test_bit_of_bl rev_nth nth_append) 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 MinI: assumes fa: "finite A" and ne: "A \ {}" and xv: "m \ A" and min: "\y \ A. m \ y" shows "Min A = m" using fa ne xv min proof (induct A arbitrary: m rule: finite_ne_induct) case singleton then show ?case by simp next case (insert y F) from insert.prems have yx: "m \ y" and fx: "\y \ F. m \ y" by auto have "m \ insert y F" by fact then show ?case proof assume mv: "m = y" have mlt: "m \ Min F" by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mv [symmetric]) apply (auto simp: min_def mlt) done next assume "m \ F" then have mf: "Min F = m" by (rule insert.hyps(4) [OF _ fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mf) apply (rule iffD2 [OF _ yx]) apply (auto simp: min_def) done qed qed 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 not_empty_eq: "(S \ {}) = (\x. x \ S)" by auto lemma range_subset_lower: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ c \ a" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_upper: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ b \ d" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d)" apply (insert non_empty) apply (rule iffI) apply (frule range_subset_lower [where x=a], simp) apply (drule range_subset_upper [where x=a], simp) apply simp apply auto done lemma range_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} = {c..d}) = (a = c \ b = d)" by (metis atLeastatMost_subset_iff eq_iff non_empty) lemma range_strict_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d \ (a = c \ b \ d))" apply (insert non_empty) apply (subst psubset_eq) apply (subst range_subset_eq, assumption+) apply (subst range_eq, assumption+) apply simp done lemma range_subsetI: fixes x :: "'a :: order" assumes xX: "X \ x" and yY: "y \ Y" shows "{x .. y} \ {X .. Y}" using xX yY by auto lemma set_False [simp]: "(set bs \ {False}) = (True \ set bs)" by auto declare of_nat_power [simp del] (* TODO: move to word *) 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 word_unat_power 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 is_aligned_0'[simp]: "is_aligned 0 n" by (simp add: is_aligned_def) lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp 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_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_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 add: word_unat.Rep_inject [symmetric]) apply simp done lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done lemma int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) 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 assume szv: "sz < LENGTH('a::len)" show ?thesis proof (cases "sz = 0") case True then show ?thesis using kv szv by (simp add: unat_of_nat) next case False then have sne: "0 < sz" .. have uk: "unat (of_nat k :: 'a word) = k" apply (subst unat_of_nat) apply (simp add: nat_mod_eq less_trans[OF kv] sne) done show ?thesis using szv apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: uk nat_less_power_trans[OF kv order_less_imp_le [OF szv]])+ done qed next assume "\ sz < LENGTH('a)" with kv show ?thesis by (simp add: not_less power_overflow) qed lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_def word_mod_by_0) 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 aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply (simp add: word_size_bl) apply (erule iffD1 [OF word_less_sub_le[OF szv]]) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_def power_overflow) then show ?thesis .. qed lemma less_two_pow_divD: "\ (x :: nat) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (rule context_conjI) apply (rule ccontr) apply (simp add: power_strict_increasing) apply (simp add: power_sub) done lemma less_two_pow_divI: "\ (x :: nat) < 2 ^ (n - m); m \ n \ \ x < 2 ^ n div 2 ^ m" by (simp add: power_sub) 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 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) (* shadows the slightly weaker Word.nth_ucast *) lemma nth_ucast: "(ucast (w::'a::len0 word)::'b::len0 word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" by (simp add: ucast_def test_bit_bin word_ubin.eq_norm nth_bintr word_size) (fast elim!: bin_nth_uint_imp) lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by (meson Word.nth_ucast test_bit_conj_lt le_def upper_bits_unset_is_l2p) 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) by word_eqI_solve 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 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 is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: word_unat_power) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" apply (simp add: word_less_nat_alt) apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply simp+ done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt) apply (subst (asm) unat_of_nat) apply (subst (asm) mod_less) apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp apply assumption done ultimately show ?thesis by auto qed lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" apply (erule udvd_minus_le') apply (simp add: udvd_def)+ done 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 sublist_equal_part: "prefix xs ys \ take (length xs) ys = xs" by (clarsimp simp: prefix_def) 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 prefix_length_less: "strict_prefix xs ys \ length xs < length ys" apply (clarsimp simp: strict_prefix_def) apply (frule prefix_length_le) apply (rule ccontr, simp) apply (clarsimp simp: prefix_def) done lemmas take_less = take_strict_prefix lemma not_prefix_longer: "\ length xs > length ys \ \ \ prefix xs ys" by (clarsimp dest!: prefix_length_le) 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 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 (subst unat_of_nat_eq[where x=n, symmetric], simp+) 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: unat_of_nat) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_def [symmetric]) apply (subst unat_def [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" by (drule unat_of_nat_eq, simp) 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 of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma map_prefixI: "prefix xs ys \ prefix (map f xs) (map f ys)" by (clarsimp simp: prefix_def) lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q as bs" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ \ P (x # xs) (y # ys)" shows "P as bs" proof - define as' where "as' == as" define bs' where "bs' == bs" have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp then show ?thesis using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 show ?case by fact next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" by (auto simp: as'_def bs'_def) qed qed qed 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 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 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 shiftr_less_t2n': "\ x && 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]) apply word_eqI apply (erule_tac x="na + n" in allE) apply fastforce 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) && ~~ (mask m) = 0" apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr) apply (subgoal_tac "w >> n + m = 0", simp) apply (simp add: le_mask_iff[symmetric] mask_def le_def) apply (subst (asm) p2_gt_0[symmetric]) apply (simp add: power_add not_less) done 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]) apply word_eqI apply (erule_tac x="na - n" in allE) apply auto 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 ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))" by word_eqI 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 :: len0) word) = unat x mod 2 ^ (LENGTH('a))" apply (simp add: unat_def ucast_def) apply (subst word_uint.eq_norm) apply (subst nat_mod_distrib) apply simp apply simp apply (subst nat_power_eq) apply simp apply simp done 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 sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply (simp add: scast_def) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp 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 nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) apply (simp add: nth_rev) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p || d" by (rule word_plus_and_or_coroll, word_eqI) blast lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule word_leq_le_minus_one) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) && mask n = q && mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma take_prefix: "(take (length xs) ys = xs) = prefix xs ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done lemma mask_out_sub_mask: "(x && ~~ (mask n)) = x - (x && (mask n))" by (simp add: field_simps word_plus_and_or_coroll2) lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) 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 shiftl_mask_is_0[simp]: "(x << n) && mask n = 0" apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done definition sum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a + 'c \ 'b + 'd" where "sum_map f g x \ case x of Inl v \ Inl (f v) | Inr v' \ Inr (g v')" lemma sum_map_simps[simp]: "sum_map f g (Inl v) = Inl (f v)" "sum_map f g (Inr w) = Inr (g w)" by (simp add: sum_map_def)+ lemma if_and_helper: "(If x v v') && v'' = If x (v && v'') (v' && v'')" by (rule if_distrib) 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 lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] lemma rshift_sub_mask_eq: "(a >> (size a - b)) && mask b = a >> (size a - b)" 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_def word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) && mask (size a - c)" apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m\size w \ (w && mask m) >> n = (w >> n) && mask (m-n)" by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w && mask m) << n = (w << n) && mask (m+n)" 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" by (simp add: le_mask_iff shiftl_shiftr3) lemma and_not_mask_twice: "(w && ~~ (mask n)) && ~~ (mask m) = w && ~~ (mask (max m n))" apply (simp add: and_not_mask) apply (case_tac "n 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 eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a && mask b = mask (min a b)" by word_eqI lemma mask_eq_0_eq_x: "(x && w = 0) = (x && ~~ w = x)" using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x && w = x) = (x && ~~ w = 0)" using word_plus_and_or_coroll2[where x=x and w=w] by auto definition "limited_and (x :: 'a :: len word) y = (x && y = x)" lemma limited_and_eq_0: "\ limited_and x z; y && ~~ z = y \ \ x && y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(&&)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y && z = z \ \ x && y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_def, folded limited_and_def] lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)" by simp lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] compl_of_1 shiftl_shiftr1[unfolded word_size mask_def] shiftl_shiftr2[unfolded word_size mask_def] lemma split_word_eq_on_mask: "(x = y) = (x && m = y && m \ x && ~~ m = y && ~~ m)" by safe word_eqI_solve lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n, simp) apply (case_tac n; simp) done lemma word_and_1_shiftl: "x && (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" by word_eqI_solve 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 && (mask n << m) = ((x >> m) && mask n) << m" by word_eqI_solve lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) lemma word_FF_is_mask: "0xFF = mask 8" by (simp add: mask_def) lemma word_1FF_is_mask: "0x1FF = mask 9" by (simp add: mask_def) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply (rule sym, subst word_unat.inverse_norm) apply (simp add: ucast_def word_of_int[symmetric] of_nat_nat[symmetric] unat_def[symmetric]) apply (simp add: unat_of_nat) 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 range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp 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 mod_mod_power: fixes k :: nat shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" proof (cases "m \ n") case True then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m" apply - apply (subst mod_less [where n = "2 ^ n"]) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp+ done also have "\ = k mod 2 ^ (min m n)" using True by simp finally show ?thesis . next case False then have "n < m" by simp then obtain d where md: "m = n + d" by (auto dest: less_imp_add_positive) then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n" by (simp add: mod_mult2_eq power_add) then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n" by (simp add: mod_add_left_eq) then show ?thesis using False by simp qed 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 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 nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done lemma power_mod_div: fixes x :: "nat" shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS") proof (cases "n \ m") case True then have "?LHS = 0" apply - apply (rule div_less) apply (rule order_less_le_trans [OF mod_less_divisor]; simp) done also have "\ = ?RHS" using True by simp finally show ?thesis . next case False then have lt: "m < n" by simp then obtain q where nv: "n = m + q" and "0 < q" by (auto dest: less_imp_Suc_add) then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m" by (simp add: power_add mod_mult2_eq) then have "?LHS = x div 2 ^ m mod 2 ^ q" by (simp add: div_add1_eq) also have "\ = ?RHS" using nv by simp finally show ?thesis . qed 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 (simp add: ucast_def word_of_int_nat unat_def) lemmas if_fun_split = if_apply_def2 lemma i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" - apply (safe, simp_all) + apply rule apply (rule ccontr) - apply (drule plus_one_helper2) + 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 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 If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp 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 word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] 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 add: word_less_nat_alt) apply (subst unat_of_nat_eq) apply (erule order_less_trans) apply simp+ 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 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 lemma of_bool_nth: "of_bool (x !! v) = (x >> v) && 1" by (word_eqI cong: rev_conj_cong) 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 (induct x) apply clarsimp+ by (metis Suc_eq_plus1 add_lessD1 less_irrefl one_add_one unatSuc word_less_nat_alt) 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) 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 && 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 shiftr_mask_eq: "(x >> n) && mask (size x - n) = x >> n" for x :: "'a :: len word" by word_eqI_solve lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) && mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) 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 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 (simp add: word_unat.norm_eq_iff [symmetric]) done lemma mask_AND_NOT_mask: "(w && ~~ (mask n)) && mask n = 0" by word_eqI lemma AND_NOT_mask_plus_AND_mask_eq: "(w && ~~ (mask n)) + (w && mask n) = w" by (subst word_plus_and_or_coroll; word_eqI_solve) lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x && mask n = y && mask n" and m2: "x && ~~ (mask n) = y && ~~ (mask n)" shows "x = y" proof (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x && mask n) !! m)" by (simp add: word_size test_bit_conj_lt) also have "\ = ((y && mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size test_bit_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x && ~~ (mask n)) !! m)" by (simp add: neg_mask_test_bit test_bit_conj_lt) also have "\ = ((y && ~~ (mask n)) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_test_bit test_bit_conj_lt) finally show ?thesis . qed qed lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) 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_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: unat_of_nat) done 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 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 and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x && y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply word_eqI done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x && 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = 1" by (simp add: mask_def) 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 (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done 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 bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x && 1) = (x && 1 = 1)" apply (rule iffI) prefer 2 apply simp apply (subgoal_tac "x && mask 1 < 2^1") prefer 2 apply (rule and_mask_less_size) apply (simp add: word_size) apply (simp add: mask_def) apply (drule word_less_cases [where y=2]) apply (erule disjE, simp) apply simp done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" by (simp add: scast_def ucast_def sint_eq_uint) 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 drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp done lemma word_0_sle_from_less[unfolded word_size]: "\ x < 2 ^ (size x - 1) \ \ 0 <=s x" apply (clarsimp simp: word_sle_msb_le) apply (simp add: word_msb_nth) apply (subst (asm) word_test_bit_def [symmetric]) apply (drule less_mask_eq) apply (drule_tac x="size x - 1" in word_eqD) apply (simp add: word_size) done lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp done lemma distinct_lemma: "f x \ f y \ x \ y" by auto 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" by unat_arith 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) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* 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: "\ \ 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 \ \ P" apply atomize_elim apply (case_tac "len_of TYPE ('a) = 1") apply clarsimp apply (subgoal_tac "(UNIV :: 'a word set) = {0, 1}") apply (metis UNIV_I insert_iff singletonE) apply (subst word_unat.univ) apply (clarsimp simp: unats_def image_def) apply (rule set_eqI, rule iffI) apply clarsimp apply (metis One_nat_def less_2_cases of_nat_1 semiring_1_class.of_nat_0) apply clarsimp apply (metis Abs_fnat_hom_0 Suc_1 lessI of_nat_1 zero_less_Suc) apply clarsimp apply (metis One_nat_def arith_is_1 le_def len_gt_0) done lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply clarsimp apply (metis Suc_pred int_word_uint len_gt_0 power_Suc uint_eq_0 word_of_int_2p word_pow_0) done lemma sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_def word_sless_alt sbintrunc_If) (* Basic proofs that signed word div/mod operations are * truncations of their integer counterparts. *) lemma signed_div_arith: "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: sdiv_int_def sdiv_word_def) apply (metis word_ubin.eq_norm) done lemma signed_mod_arith: "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: smod_int_def smod_word_def) apply (metis word_ubin.eq_norm) done (* Signed word arithmetic overflow constraints. *) 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)+ lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply (simp add: sint_word_ariths scast_def) apply (simp add: wi_hom_mult) apply (subst word_sint.Abs_inject, simp_all) apply (simp add: sints_def range_sbintrunc abs_less_iff) apply clarsimp apply (simp add: sints_def range_sbintrunc word_size) apply (auto elim: order_less_le_trans order_trans[rotated]) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" apply (auto simp: sdiv_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) apply (clarsimp simp: sdiv_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if) apply (clarsimp simp: algebra_split_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") apply clarsimp apply (rule classical) apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) apply (clarsimp simp: sdiv_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" apply (clarsimp simp: sdiv_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if algebra_split_simps not_less) apply (case_tac "sgn (a * b) = -1") apply (clarsimp simp: not_less algebra_split_simps) apply (clarsimp simp: algebra_split_simps not_less) apply (rule classical) apply (case_tac "b = 0") apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" apply (unfold sdiv_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (clarsimp simp: sgn_if) apply (meson abs_ge_zero neg_le_0_iff_le nonneg_mod_div order_trans) apply (metis abs_eq_0 abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) apply (clarsimp simp: sdiv_word_def sdiv_int_def) apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" by (clarsimp simp: sdiv_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" by (clarsimp simp: sdiv_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) apply (metis wi_hom_neg word_sint.Rep_inverse') done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done 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: sdiv_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] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where a="numeral x" for x] sdiv_word_def[where a=0] sdiv_word_def[where a=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where b="numeral y" for y] word_sdiv_numerals_lhs[where b=0] word_sdiv_numerals_lhs[where b=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" apply (clarsimp simp: smod_int_def sdiv_int_def) apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] apply (auto simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute])[1] apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" by (clarsimp simp: smod_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (case_tac "a \ - (2 ^ (LENGTH('a) - 1)) \ b \ -1") apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def minus_word.abs_eq [symmetric] times_word.abs_eq [symmetric]) apply (clarsimp simp: smod_word_def smod_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where a="numeral x" for x] smod_word_def[where a=0] smod_word_def[where a=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where b="numeral y" for y] word_smod_numerals_lhs[where b=0] word_smod_numerals_lhs[where b=1] lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" apply (clarsimp simp: word_of_int int_word_sint) apply (subst int_mod_eq') apply simp apply (subst (2) power_minus_simp) apply clarsimp apply clarsimp apply clarsimp done lemma of_int_sint [simp]: "of_int (sint a) = a" apply (insert word_sint.Rep [where x=a]) apply (clarsimp simp: word_of_int) done lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" apply (subst nth_w2p) apply (case_tac "n \ LENGTH('a)") apply (subst power_overflow, simp) apply clarsimp apply (metis nth_w2p scast_def test_bit_conj_lt len_signed nth_word_of_int word_sint.Rep_inverse) done lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) 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 (clarsimp simp: word_of_int ucast_def) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_of_int) apply (subst word.abs_eq_iff) 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_mod_power_int 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_def 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_def 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_def 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_def 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_def 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_def 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_def 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_def 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 (clarsimp simp: scast_def) apply (metis down_cast_same is_up_down scast_def 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_bl 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 smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done 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 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 (subst sint_eq_uint) apply (clarsimp simp: msb_nth nth_ucast is_down) apply (metis Suc_leI Suc_pred len_gt_0) apply (clarsimp simp: uint_up_ucast is_up is_down) done lemma word_less_nowrapI': "(x :: 'a :: len0 word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = 2 ^ n" by (clarsimp simp: mask_def) 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_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp add: ucast_nat_def [symmetric]) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) 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 add: ucast_nat_def [symmetric]) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" apply (subgoal_tac "\ msb (ucast b :: 'b signed word)") apply (clarsimp simp: word_sle_msb_le) apply (clarsimp simp: is_down not_le msb_nth nth_ucast) apply (subst (asm) test_bit_conj_lt [symmetric]) apply clarsimp apply arith done lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (LENGTH('a) - 1))" apply (case_tac "msb b") apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) done (* to_bool / from_bool. *) definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x && 1) = (x !! 0)" apply (simp add: to_bool_def) apply (rule iffI) apply (rule classical, erule notE, word_eqI) apply (case_tac n; simp) apply (rule notI, drule word_eqD[where x=0]) apply simp done lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) && 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast nth_shiftr nth_rev field_simps) apply (simp add: length_word_rsplit_exp_size) apply (metis mult.commute given_quot_alt word_size word_size_gt_0) done lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma neg_mask_add_mask: "((x:: 'a :: len word) && ~~ (mask n)) + (2 ^ n - 1) = x || mask n" unfolding mask_2pm1[symmetric] by (subst word_plus_and_or_coroll; word_eqI_solve) lemma subtract_mask: "p - (p && mask n) = (p && ~~ (mask n))" "p - (p && ~~ (mask n)) = (p && mask n)" by (simp add: field_simps word_plus_and_or_coroll2)+ lemma and_neg_mask_plus_mask_mono: "(p && ~~ (mask n)) + mask n \ p" apply (rule word_le_minus_cancel[where x = "p && ~~ (mask n)"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_def word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr && ~~ (mask n)) + (2 ^ n - 1)" by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply simp done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule word_le_minus_one_leq) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d && mask n = d) \ (p + d && (~~ (mask n)) = p)" apply (subst(asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll; word_eqI; blast) using word_plus_and_or_coroll2[where x="p + d" and w="mask n"] by simp lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p && mask n = d) \ (p && (~~ (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma mask_twice: "(x && mask n) && mask m = x && mask (min m n)" by word_eqI_solve lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k && mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a && mask n = (ptr && mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a && mask n) && ~~ (mask m) = (ptr + a && ~~ (mask m) ) && mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) 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 && mask n) && ~~ (mask n) = 0" apply (clarsimp simp:mask_def) by (metis word_bool_alg.conj_cancel_right word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma and_and_not[simp]:"(a && b) && ~~ b = 0" apply (subst word_bw_assocs(1)) apply clarsimp done lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0" apply (clarsimp simp:mask_def) by (metis (erased, hide_lams) mask_eq_x_eq_0 shiftl_over_and_dist word_bool_alg.conj_absorb word_bw_assocs(1)) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma ex_mask_1[simp]: "(\x. mask x = 1)" apply (rule_tac x=1 in exI) apply (simp add:mask_def) done lemma not_switch:"~~ a = x \ a = ~~ x" by auto (* 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 && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)" by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\~~ a = b << c; \x. b = mask x\ \ (x && a || (y && b << c)) && a = x && a" apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_def) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y" by (metis (mono_tags) min_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma bit_twiddle_max: "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y" by (metis (mono_tags) max_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma swap_with_xor: "\(x::'a::len word) = a xor b; y = b xor x; z = x xor y\ \ z = b \ y = a" by (metis word_bool_alg.xor_assoc word_bool_alg.xor_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right) lemma scast_nop1: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop2: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop[simp] = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x && mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) || ~~ (mask n)) && mask n = x && 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) || y && mask n) && ~~ (mask m) = x && ~~ (mask m)" apply (subst word_ao_dist) apply (subgoal_tac "(y && mask n) && ~~ (mask m) = 0") apply simp by (metis (no_types, hide_lams) is_aligned_mask is_aligned_weaken word_and_not word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w && ~~(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) && mask m) && mask n = x && mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" apply clarsimp apply (subgoal_tac "2 \ uints (LENGTH('a))") apply (subst (asm) Word.word_ubin.set_iff_norm) apply simp apply (subst word_uint.set_iff_norm) apply clarsimp apply (rule int_mod_eq') apply simp apply (rule pow_2_gt) apply simp done 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" proof (cases \2 \ LENGTH('a)\) case False then have *: \LENGTH('a) = 1\ by simp then have \x = 0 \ x = 1\ by (metis One_nat_def less_irrefl_nat sint_1_cases) then show ?thesis by (auto simp add: word_arith_nat_defs(6) *) next case True then show ?thesis using shiftr1_unfold [symmetric, of x] uint_2_id [where ?'a = 'a] by (simp add: shiftr1_def word_div_def) qed 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 (case_tac "n = 0") apply clarsimp apply (subst div_by_0_word) apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) apply (metis (poly_guards_query) One_nat_def less_one nat_neq_iff unat_eq_1(2) unat_eq_zero) 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 max_word_max plus_1_less) + 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" by (metis add.commute add_left_cancel max_word_max not_less word_and_1 word_bool_alg.conj_one_right word_bw_comms(1) word_overflow zero_neq_one) lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" unfolding word_lsb_def bin_last_def by (metis (no_types, hide_lams) nat_mod_distrib nat_numeral not_mod_2_eq_1_eq_0 numeral_One uint_eq_0 uint_nonnegative unat_0 unat_def zero_le_numeral) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" apply (simp add:even_iff_mod_2_eq_zero) apply (subst word_lsb_nat[unfolded One_nat_def, symmetric]) apply (rule word_lsb_alt) done 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 (case_tac "LENGTH('a) = 1") apply clarsimp apply (drule_tac x=x in degenerate_word[unfolded One_nat_def]) apply (erule disjE) apply clarsimp+ apply (subst (asm) shiftr1_is_div_2[unfolded One_nat_def])+ apply (subst (asm) word_arith_nat_div)+ apply clarsimp apply (subst (asm) bintrunc_id) apply (subgoal_tac "LENGTH('a) > 0") apply linarith apply clarsimp+ apply (subst (asm) bintrunc_id) apply (subgoal_tac "LENGTH('a) > 0") apply linarith apply clarsimp+ apply (case_tac "x + 1 = 0") apply (clarsimp simp:overflow_imp_lsb) apply (cut_tac x=x in word_overflow_unat) apply clarsimp apply (case_tac "even (unat x)") apply (subgoal_tac "unat x div 2 = Suc (unat x) div 2") apply metis apply (subst numeral_2_eq_2)+ apply simp apply (simp add:odd_iff_lsb) 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) || (((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 (clarsimp simp:mask_def) - apply (metis max_word_eq max_word_max mult_2_right) - apply (metis add_diff_cancel_left' diff_less len_gt_0 mult_2_right) + apply (simp_all add: mask_def 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) || (((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 && mask n || x && ~~ (mask n) = x" 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) && ~~ (mask n) = p && ~~ (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 (cut_tac n=x and 'a='a in max_word_max, - clarsimp simp:max_word_def, - simp add: less_le) + 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 (max_word::'b::len word)" - shows "ucast ((ucast x)::'b word) = x" + \UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ + if \x \ UCAST('b::len \ 'a) (- 1)\ + for x :: \'a::len word\ proof - - { - assume a1: "x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1)::'b word))" - 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) - } - with assms show ?thesis - unfolding ucast_def - apply (subst word_ubin.eq_norm) - apply (subst and_mask_bintr[symmetric]) - apply (subst and_mask_eq_iff_le_mask) - apply (clarsimp simp: max_word_def mask_def) - done + from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ + by (simp add: ucast_def word_of_int_minus) + 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 + by (simp add: ucast_def word_ubin.eq_norm bintrunc_mod2p) 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 unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_def 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 word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral) lemma complement_nth_w2p: shows "n' < LENGTH('a) \ (~~ (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 && 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) && 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 td_gal_lt) lemma le_or_mask: "w \ w' \ w || mask x \ w' || mask x" 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 (unfold word_le_def shiftr1_def word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (case_tac "uint v" rule: bin_exhaust) by (smt Bit_def bin_rest_BIT) lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len0 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_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def - by (simp add: max_word_bl[simplified max_word_minus] takeWhile_replicate) + by (simp add: takeWhile_replicate) lemma word_add_no_overflow:"(x::'a::len word) < max_word \ 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_def unat_def 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 add: ucast_nat_def[symmetric]) 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] (* 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 Word.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 minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) && ~~ (mask n) = 0" by (simp add: and_not_mask shiftr_eq_0) lemma two_power_strict_part_mono: "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" proof - { fix n have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" proof (induct n) case 0 then show ?case by simp next case (Suc n) from Suc.prems have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce with Suc show ?case by (subst strict_part_mono_by_steps) simp qed } then show ?thesis by simp qed 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_sub_int) 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 (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp 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) unfolding is_aligned_def by (metis le_unat_uoi nat_dvd_not_less order_less_imp_le unat_power_lower) 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_def by (metis max_word_eq shiftl_1) + unfolding mask_def by simp lemmas mask_len_max = max_word_mask[symmetric] lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def complement_def is_aligned_mask mask_def word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma complement_mask: "complement (2 ^ n - 1) = ~~ (mask n)" unfolding complement_def mask_def by simp lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis complement_mask is_aligned_add_helper p_assoc_help power_2_ge_iff) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans unat_less_power word_less_sub_le word_mod_less_divisor) have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst complement_mask) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (metis unat_div unat_less_helper unat_power_lower) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff unat_def zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 && ~~ (mask sz)" unfolding alignUp_def[unfolded complement_def] by (simp add:mask_def[symmetric,unfolded shiftl_t2n,simplified]) lemma mask_out_first_mask_some: "\ x && ~~ (mask n) = y; n \ m \ \ x && ~~ (mask m) = y && ~~ (mask m)" by word_eqI_solve lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" by (simp add: aligned_add_offset_less) lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q && ~~ (mask n)) = (p + q) && ~~ (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 && ~~ (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma mask_lower_twice: "n \ m \ (x && ~~ (mask n)) && ~~ (mask m) = x && ~~ (mask m)" by word_eqI_solve lemma mask_lower_twice2: "(a && ~~ (mask n)) && ~~ (mask m) = a && ~~ (mask (max n m))" by word_eqI_solve lemma ucast_and_neg_mask: "ucast (x && ~~ (mask n)) = ucast x && ~~ (mask n)" by word_eqI_solve lemma ucast_and_mask: "ucast (x && mask n) = ucast x && mask n" by word_eqI_solve lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x && mask n) :: 'a word) = ucast x" by word_eqI lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q && ~~ (mask sz)) = (p - ((alignUp q sz) && ~~ (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (rule sum_to_zero) apply (subst add.commute) by (simp add: alignUp_distance and_mask_0_iff_le_mask is_aligned_neg_mask_eq mask_out_add_aligned) lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by (simp add: is_down ucast_id ucast_ucast_a) 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 word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y" unfolding ucast_def Word.bitOR_word.abs_eq uint_or by blast 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 && w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed 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 from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast 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]: - "max_word \ 0" - by (simp add: max_word_minus) +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 && mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed -lemma unat_max_word_pos[simp]: "0 < unat max_word" - by (auto simp: unat_gt_0) +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 (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) 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]) (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_test_bit word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_if'[word_eqI_simps] = sign_extend_bitwise_if[simplified word_size] lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w || ~~ (mask (Suc n)) else w && mask (Suc n))" by word_eqI (auto dest: less_antisym) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if') lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply (rule iffI) apply (word_eqI, rename_tac i) apply (case_tac "n < i"; simp add: sign_extended_def word_size) apply (erule subst, rule sign_extended_sign_extend) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by word_eqI lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w && mask (Suc n) = v && mask (Suc n) \ sign_extend n w = sign_extend n v" by word_eqI_solve lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f && mask e = f" by (fastforce intro: subst[where P="\f. f && mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr && ~~ (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) (* 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) by (meson less_irrefl not_less zle2p) lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "w BIT c \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num Bit_def) lemma Bit_in_uintsI: "w BIT c \ 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" using that proof (induction n arbitrary: b m) case 0 then have "uints l \ uints m" by (intro uints_monoI) auto then show ?case using 0 by auto next case (Suc n) then show ?case by (auto intro!: Bit_in_uintsI) 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) proof (induction m arbitrary: b d) case (Suc m) show ?case using Suc.prems by (auto intro: Suc.IH) qed simp lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (induction n arbitrary: b d) auto lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (induction n arbitrary: b d) auto 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::len0 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::len0 word" and b::"'b::len0 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::len0 word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len0 word" and b::"'b::len0 word" by (auto simp: word_cat_def word_uint.Abs_inject word_of_int_bin_cat_eq_iff that) 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 unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) 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_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)" by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len word_ctz_le) 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^n + mask n" by (simp add: mask_def) lemma is_aligned_no_overflow_mask: "is_aligned x n \ x \ x + mask n" by (simp add: mask_def) (erule is_aligned_no_overflow') lemma is_aligned_mask_offset_unat: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off \ mask sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) from offv szv have offv': "unat off < 2 ^ sz" by (simp add: mask_2pm1 unat_less_power) show ?thesis using szv using al is_aligned_no_wrap''' offv' by blast next assume "\ sz < LENGTH('a)" with al have "x = 0" by - word_eqI thus ?thesis by simp qed lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" apply (induct xs) apply simp apply (simp add: of_bl_Cons mask_Suc) apply (rule conjI; clarsimp) apply (erule word_plus_mono_right) apply (rule is_aligned_no_overflow_mask) apply (rule is_aligned_triv) apply (simp add: word_le_nat_alt) apply (subst unat_add_lem') apply (rule is_aligned_mask_offset_unat) apply (rule is_aligned_triv) apply (simp add: mask_def) apply simp done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_def) lemma is_aligned_over_length: "\ is_aligned p n; LENGTH('a) \ n \ \ (p::'a::len word) = 0" by (simp add: is_aligned_mask mask_over_length) 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 is_aligned_add_step_le: "\ is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \ a + mask n \ \ False" apply (simp flip: not_le) apply (erule notE) apply (cases "LENGTH('a) \ n") apply (drule (1) is_aligned_over_length)+ apply (drule mask_over_length) apply clarsimp apply (clarsimp simp: word_le_nat_alt not_less) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def word_le_nat_alt) apply (drule le_imp_less_Suc) apply (simp add: Suc_2p_unat_mask) by (metis Groups.mult_ac(2) Suc_leI linorder_not_less mult_le_mono order_refl times_nat.simps(2)) lemma power_2_mult_step_le: "\n' \ n; 2 ^ n' * k' < 2 ^ n * k\ \ 2 ^ n' * (k' + 1) \ 2 ^ n * (k::nat)" apply (cases "n'=n", simp) apply (metis Suc_leI le_refl mult_Suc_right mult_le_mono semiring_normalization_rules(7)) apply (drule (1) le_neq_trans) apply clarsimp apply (subgoal_tac "\m. n = n' + m") prefer 2 apply (simp add: le_Suc_ex) apply (clarsimp simp: power_add) by (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj) lemma aligned_mask_step: "\ n' \ n; p' \ p + mask n; is_aligned p n; is_aligned p' n' \ \ (p'::'a::len word) + mask n' \ p + mask n" apply (cases "LENGTH('a) \ n") apply (frule (1) is_aligned_over_length) apply (drule mask_over_length) apply clarsimp apply (simp add: not_le) apply (simp add: word_le_nat_alt unat_plus_simple) apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+ apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def) apply (rename_tac k k') apply (thin_tac "unat p = x" for p x)+ apply (subst Suc_le_mono[symmetric]) apply (simp only: Suc_2p_unat_mask) apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption) apply (erule (1) power_2_mult_step_le) done 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 && b = 0" by word_eqI_solve lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a || b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemmas word_and_or_mask_aligned2 = word_and_or_mask_aligned[where a=b and b=a for a b, simplified add.commute word_bool_alg.disj.commute] lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" by (clarsimp simp: word_eqI_simps) 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 is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned2) 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::len0 word) \ mask n" by (clarsimp simp: le_mask_high_bits word_size nth_ucast) 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 aligned_add_mask_lessD: "\ x + mask n < y; is_aligned x n \ \ x < y" for y::"'a::len word" by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans) lemma aligned_add_mask_less_eq: "\ is_aligned x n; is_aligned y n; n < LENGTH('a) \ \ (x + mask n < y) = (x < y)" for y::"'a::len word" using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast 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 fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) 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: "m && (2 ^ n - 1) = 0 \ (x + y * m) && ~~(mask n) = (x && ~~(mask n)) + y * m" by (metis Groups.add_ac(2) is_aligned_mask mask_def mask_eqs(5) mask_out_add_aligned mult_zero_right shiftl_1 word_bw_comms(1) word_log_esimps(1)) 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: 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 && mask n = 0 \ x + y && ~~(mask n) = (x && ~~(mask n)) + y" 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 && mask n = 0; y && 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_def) done lemma sub_right_shift: "\ x && mask n = 0; y && 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 && mask n = mask n \ (x && y) && mask n = x && mask n" by (simp add: word_bool_alg.conj.assoc) lemma and_and_mask_simple_not: "y && mask n = 0 \ (x && y) && mask n = 0" by (simp add: word_bool_alg.conj.assoc) lemma word_and_le': "b \ c \ (a :: 'a :: len word) && b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) && b < c" by (metis word_and_le1 xtr7) 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 ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x && 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 && mask n = 0; m \ n \ \ x && mask m = 0" by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) && 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) && 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma pow_mono_leq_imp_lt: "x \ y \ x < 2 ^ y" by (simp add: le_less_trans) lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by simp lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len sword) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by (metis le_unat_uoi le_unat_uoi linorder_neqE_nat nat_less_le scast_of_nat word_unat.Rep_inverse) lemma shiftr_and_eq_shiftl: "(w >> n) && x = y \ w && (x << n) = (y << n)" for y :: "'a:: len word" by (smt and_not_mask is_aligned_neg_mask_eq is_aligned_shift shift_over_ao_dists(4) word_bool_alg.conj_assoc word_bw_comms(1)) lemma neg_mask_combine: "~~(mask a) && ~~(mask b) = ~~(mask (max a b))" by (auto simp: word_ops_nth_size word_size intro!: word_eqI) lemma neg_mask_twice: "x && ~~(mask n) && ~~(mask m) = x && ~~(mask (max n m))" by (metis neg_mask_combine) lemma multiple_mask_trivia: "n \ m \ (x && ~~(mask n)) + (x && mask n && ~~(mask m)) = x && ~~(mask m)" 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 add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p && ~~(mask n) = x" using add_mask_lower_bits by auto lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' && ~~(mask bits) = ptr) = (ptr' \ mask_range ptr bits)" apply (erule is_aligned_get_word_bits) apply (rule iffI) apply (drule sym) apply (simp add: word_and_le2) apply (subst word_plus_and_or_coroll, word_eqI_solve) apply (metis le_word_or2 neg_mask_add_mask word_bool_alg.conj.right_idem) apply clarsimp apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le word_and_not) apply (simp add: power_overflow mask_def) done lemma aligned_offset_in_range: "\ is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \ m; n < LENGTH('a) \ \ (x + y \ {p .. p + mask n}) = (x \ mask_range p n)" apply (simp only: is_aligned_add_or flip: neg_mask_in_mask_range) by (metis less_mask_eq mask_subsume) lemma mask_range_to_bl': "\ is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \ \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (rule set_eqI, rule iffI) apply clarsimp apply (subgoal_tac "\y. x = ptr + y \ y < 2 ^ bits") apply clarsimp apply (subst is_aligned_add_conv) apply assumption apply simp apply simp apply (rule_tac x="x - ptr" in exI) apply (simp add: add_diff_eq[symmetric]) apply (simp only: word_less_sub_le[symmetric]) apply (rule word_diff_ls') apply (simp add: field_simps mask_def) apply assumption apply simp apply (subgoal_tac "\y. y < 2 ^ bits \ to_bl (ptr + y) = to_bl x") apply clarsimp apply (rule conjI) apply (erule(1) is_aligned_no_wrap') apply (simp only: add_diff_eq[symmetric] mask_def) apply (rule word_plus_mono_right) apply simp apply (erule is_aligned_no_wrap') apply simp apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI) apply (rule context_conjI) apply (rule order_less_le_trans [OF of_bl_length]) apply simp apply simp apply (subst is_aligned_add_conv) apply assumption apply simp apply (drule sym) apply (simp add: word_rep_drop) done lemma mask_range_to_bl: "is_aligned (ptr :: 'a :: len word) bits \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (erule is_aligned_get_word_bits) apply (erule(1) mask_range_to_bl') apply (rule set_eqI) apply (simp add: power_overflow mask_def) done lemma aligned_mask_range_cases: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \ \ mask_range p n \ mask_range p' n' = {} \ mask_range p n \ mask_range p' n' \ mask_range p n \ mask_range p' n'" apply (simp add: mask_range_to_bl) apply (rule Meson.disj_comm, rule disjCI) apply (erule nonemptyE) apply simp apply (subgoal_tac "(\n''. LENGTH('a) - n = (LENGTH('a) - n') + n'') \ (\n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')") apply (fastforce simp: take_add) apply arith done lemma aligned_mask_range_offset_subset: assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" shows "mask_range (ptr+x) sz' \ mask_range ptr sz" using al proof (rule is_aligned_get_word_bits) assume p0: "ptr = 0" and szv': "LENGTH ('a) \ sz" then have "(2 ::'a word) ^ sz = 0" by simp show ?thesis using p0 by (simp add: \2 ^ sz = 0\ mask_def) next assume szv': "sz < LENGTH('a)" hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)" using szv by auto show ?thesis using szv szv' apply (intro range_subsetI) apply (rule is_aligned_no_wrap' [OF al xsz]) apply (simp only: flip: add_diff_eq add_mask_fold) apply (subst add.assoc, rule word_plus_mono_right) using al' is_aligned_add_less_t2n xsz apply fastforce apply (simp add: field_simps szv al is_aligned_no_overflow) done qed lemma aligned_mask_diff: "\ is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz; bits \ sz; sz < LENGTH('a); dest < ptr \ \ mask bits + dest < ptr" apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption) apply (elim disjE) apply (drule_tac is_aligned_no_overflow_mask, simp)+ apply (simp add: algebra_split_simps word_le_not_less) apply (drule is_aligned_no_overflow_mask; fastforce) by (simp add: aligned_add_mask_less_eq is_aligned_weaken algebra_split_simps) lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p && ~~(mask n') \ p'; p' && ~~(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x && mask n >> m" in spec) apply (clarsimp simp: shiftr_less_t2n and_mask_less_size wsst_TYs multiple_mask_trivia word_bw_assocs neg_mask_twice max_absorb2 shiftr_shiftl1) done 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)" by (clarsimp simp: word_le_def uint_nat of_nat_inverse) 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 && ~~(mask sz)) + 2 ^ sz - 1" 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))" apply (case_tac "LENGTH('a) = 1") apply (clarsimp simp: ucast_def) apply (metis (hide_lams, mono_tags) One_nat_def len_signed plus_word.abs_eq uint_word_arith_bintrs(1) word_ubin.Abs_norm) apply (clarsimp simp: ucast_def) apply (metis le_refl len_signed plus_word.abs_eq uint_word_arith_bintrs(1) wi_bintr) done 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) && (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) && mask m) = unat (x && mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma small_powers_of_2: "x \ 3 \ x < 2 ^ (x - 1)" by (induct x; simp add: suc_le_pow_2) lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using small_powers_of_2 by (smt One_nat_def diff_less le_less_trans len_gt_0 len_signed lessI n_less_equal_power_2 not_msb_from_less of_nat_mono sint_eq_uint uint_nat unat_of_nat_eq unat_power_lower word_clz_max word_of_nat_less wsst_TYs(3)) lemma word_clz_sint_lower[simp]: "LENGTH('a) \ 3 \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using small_powers_of_2 uint_nat apply (simp add: order_le_less_trans[OF word_clz_max] not_msb_from_less word_of_nat_less word_size) by (simp add: uint_nat) 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_def power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" using less_not_refl3 le_step_down_nat le_trans less_or_eq_imp_le word_shiftr_lt by (metis (no_types, lifting)) 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 (simp add: ucast_def word_of_int) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n && msk = mask n \ \ (x mod m) && msk = x mod m" by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x && 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 && ~~(mask n)" by (metis (no_types) AND_NOT_mask_plus_AND_mask_eq add_right_imp_eq diff_add_cancel and_mask_eq_iff_shiftr_0 shiftr_mask_le word_bool_alg.conj.commute) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr && ~~(mask sz')) - (ptr && ~~(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") proof - assume lt: "sz' \ sz" hence "?lhs = ptr && (mask sz && ~~(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_range_subsetD: "\ p' \ mask_range p n; x' \ mask_range p' n'; n' \ n; is_aligned p n; is_aligned p' n' \ \ x' \ mask_range p n" using aligned_mask_step by fastforce lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2 word_plus_mono_right) lemma of_bl_length2: "length xs + c < LENGTH('a) \ of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) && ~~(mask sz) = 0" by (simp add: Word_Lemmas.of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr && ~~(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) && ~~(mask sz)) + unat (ptr && 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) Word_Lemmas.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 && mask n) = unat d \ unat (p + d && ~~(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) && ~~(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) && mask high_bits = x >> low_bits" 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 && mask low_bits = 0) = (x = 0)" 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) && mask n = q && 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 from_to_bool_last_bit: "from_bool (to_bool (x && 1)) = x && 1" by (metis from_bool_to_bool_iff word_and_1) 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) && mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x && ~~(mask sz)) + (x && mask sz) \ (x && ~~(mask sz)) + ((x && 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 Word_Lemmas.of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "~~(mask n) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * ~~(mask n) = - (x && ~~(mask n))" 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) && ~~(mask n) = p + (x && ~~(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) && ~~(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) && mask (m::nat) = mask (min m n)" 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 && mask m = (if n < m then 2 ^ n else 0)" 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) \ \ (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 && mask a >> b) :: 'a :: len word) = p && 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 && mask a >> b) :: 'a :: len word) = unat (p && mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p && mask a) && ~~(mask b) = 0" by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p && mask a << c) && ~~(mask b) = 0" by (metis and_not_mask le_mask_iff mask_shiftl_decompose shiftl_0 shiftl_over_and_dist shiftr_mask_le word_and_le' word_bool_alg.conj_commute) lemma mask_overlap_zero': "a \ b \ (p && ~~(mask a)) && mask b = 0" 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 && ~~(mask a)) + ((p && mask a) && ~~(mask b)) + (p && mask b) = p" by (simp add: add.commute multiple_mask_trivia word_bool_alg.conj_commute word_bool_alg.conj_left_commute word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p && mask a >> b << b) = (p && mask a) && ~~(mask b)" by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p && mask b) \ \ (p && ~~(mask a)) + (p && 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 (metis scast_def word_of_int) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by (metis cast_simps(23) scast_scast_id(2)) 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 && mask n \ of_nat ` set [0 ..< 2 ^ n]" proof - have "x && mask n \ {0 .. 2 ^ n - 1}" by (simp add: mask_def word_and_le1) also have "... = of_nat ` {0 .. 2 ^ n - 1}" apply (rule set_eqI, rule iffI) apply (clarsimp simp: image_iff) apply (rule_tac x="unat x" in bexI; simp) using len apply (simp add: word_le_nat_alt unat_2tp_if unat_minus_one) using len apply (clarsimp simp: word_le_nat_alt unat_2tp_if unat_minus_one) apply (subst unat_of_nat_eq; simp add: nat_le_Suc_less) apply (erule less_le_trans) apply simp done also have "{0::nat .. 2^n - 1} = set [0 ..< 2^n]" by (auto simp: nat_le_Suc_less) finally show ?thesis . qed lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: Word_Lemmas.of_nat_power not_msb_from_less sint_eq_uint) 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 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)" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_power_minus_less of_nat_le_iff sint_eq_uint_2pl uint_nat unat_of_nat_len) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_less_le sint_eq_uint_2pl uint_nat unat_lt2p unat_of_nat_len unat_power_lower) lemma sint_ctz: "LENGTH('a) > 2 \ 0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))" apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)") apply (rule conjI) apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le) apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le) by (rule small_powers_of_2, simp) 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 && mask (size w - n)" by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) \ unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n" by (metis less_trans unat_of_nat_eq word_log2_max word_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_def NOT_eq flip: mul_not_mask_eq_neg_shiftl) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x && mask LENGTH('b) = y && 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_def 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))" unfolding scast_def sint_uint word_size apply (subst word_eq_iff) apply (rule iffI) apply (rule ballI) apply (drule_tac x=i in spec) apply (subst (asm) test_bit_sbintrunc_ucast; simp) apply (rule allI) apply (case_tac "n < LENGTH('a)") apply (subst test_bit_sbintrunc_ucast) apply simp apply (case_tac "n \ LENGTH('b)") apply (drule_tac x=n in bspec) by auto lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ ~~(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 && mask LENGTH('b))" by word_eqI_solve lemma ucast_NOT: "ucast (~~x) = ~~(ucast x) && 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) (~~x) = ~~(UCAST('a \ 'b) x)" by word_eqI 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) && ~~(mask n) = a" by (smt add.left_neutral add_diff_cancel_right' add_mask_lower_bits and_mask_plus is_aligned_mask is_aligned_weaken le_less_trans of_bl_length2 subtract_mask(1)) 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 (subst word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule is_aligned_AND_less_0) apply (simp add: is_aligned_mask) apply (rule order_less_le_trans) apply (rule of_bl_length2) apply simp apply (simp add: two_power_increasing) apply simp apply (rule nth_equalityI) apply (simp only: len_bin_to_bl) apply (clarsimp simp only: len_bin_to_bl nth_bin_to_bl word_test_bit_def[symmetric]) apply (simp add: nth_shiftr nth_shiftl shiftl_t2n[where n=c, simplified mult.commute, simplified, symmetric]) apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl nth_rev) apply arith done end diff --git a/thys/Word_Lib/Word_Lib.thy b/thys/Word_Lib/Word_Lib.thy --- a/thys/Word_Lib/Word_Lib.thy +++ b/thys/Word_Lib/Word_Lib.thy @@ -1,701 +1,694 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Additional Word Operations" theory Word_Lib imports Word_Syntax begin definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition complement :: "'a :: len word \ 'a word" where "complement x \ ~~ x" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 && complement (2 ^ n - 1)" (* standard notation for blocks of 2^n-1 words, usually aligned; abbreviation so it simplifies directly *) abbreviation mask_range :: "'a::len word \ nat \ 'a word set" where "mask_range p n \ {p .. p + mask n}" (* Haskellish names/syntax *) notation (input) test_bit ("testBit") definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" (* * Signed division: when the result of a division is negative, * we will round towards zero instead of towards minus infinity. *) class signed_div = fixes sdiv :: "'a \ 'a \ 'a" (infixl "sdiv" 70) fixes smod :: "'a \ 'a \ 'a" (infixl "smod" 70) instantiation int :: signed_div begin definition "(a :: int) sdiv b \ sgn (a * b) * (abs a div abs b)" definition "(a :: int) smod b \ a - (a sdiv b) * b" instance .. end instantiation word :: (len) signed_div begin definition "(a :: ('a::len) word) sdiv b = word_of_int (sint a sdiv sint b)" definition "(a :: ('a::len) word) smod b = word_of_int (sint a smod sint b)" instance .. end (* Tests *) lemma "( 4 :: word32) sdiv 4 = 1" "(-4 :: word32) sdiv 4 = -1" "(-3 :: word32) sdiv 4 = 0" "( 3 :: word32) sdiv -4 = 0" "(-3 :: word32) sdiv -4 = 0" "(-5 :: word32) sdiv -4 = 1" "( 5 :: word32) sdiv -4 = -1" by (simp_all add: sdiv_word_def sdiv_int_def) lemma "( 4 :: word32) smod 4 = 0" "( 3 :: word32) smod 4 = 3" "(-3 :: word32) smod 4 = -3" "( 3 :: word32) smod -4 = 3" "(-3 :: word32) smod -4 = -3" "(-5 :: word32) smod -4 = -1" "( 5 :: word32) smod -4 = 1" by (simp_all add: smod_word_def smod_int_def sdiv_int_def) (* Count leading zeros *) definition word_clz :: "'a::len word \ nat" where "word_clz w \ length (takeWhile Not (to_bl w))" (* Count trailing zeros *) definition word_ctz :: "'a::len word \ nat" where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" definition word_log2 :: "'a::len word \ nat" where "word_log2 (w::'a::len word) \ size w - 1 - word_clz w" (* Bit population count. Equivalent of __builtin_popcount. *) definition pop_count :: "('a::len) word \ nat" where "pop_count w \ length (filter id (to_bl w))" (* Sign extension from bit n *) definition sign_extend :: "nat \ 'a::len word \ 'a word" where "sign_extend n w \ if w !! n then w || ~~ (mask n) else w && mask n" definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ w !! i = w !! n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" apply (induct x) apply simp apply (simp add: shiftl1_2t) done 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 (simp add: of_bl_def) apply (rule word_uint.Abs_inverse) apply (simp add: uints_num bl_to_bin_ge0) apply (rule order_less_le_trans) apply (rule bl_to_bin_lt2p_drop) 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_ops_nth [simp]: shows word_or_nth: "(x || y) !! n = (x !! n \ y !! n)" and word_and_nth: "(x && 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])+ (* simp del to avoid warning on the simp add in iff *) declare test_bit_1 [simp del, iff] (* test: *) lemma "1 < (1024::32 word) \ 1 \ (1024::32 word)" by simp lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" 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)" apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma AND_twice [simp]: "(w && m) && m = w && m" by (simp add: word_eqI) lemma word_combine_masks: "w && m = z \ w && m' = z' \ w && (m || m') = (z || z')" by (auto simp: word_eq_iff) lemma nth_w2p_same: "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" by (simp add : nth_w2p) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" apply (simp add : word_gt_0) apply safe apply (erule swap) apply (rule word_eqI) apply (simp add : nth_w2p) apply (drule word_eqD) apply simp apply (erule notE) apply (erule nth_w2p_same [THEN iffD2]) done lemmas uint_2p_alt = uint_2p [unfolded p2_gt_0] 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 (rule word_uint.Rep_inverse' [THEN sym]) apply (rule shiftr_div_2n) done lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] lemmas p2_eq_0 [simp] = trans [OF eq_commute iffD2 [OF Not_eq_iff p2_gt_0, folded le_def, unfolded word_gt_0 not_not]] lemma neg_mask_is_div': "n < size w \ w AND NOT (mask n) = ((w div (2 ^ n)) * (2 ^ n))" by (simp add : and_not_mask shiftr_div_2n_w shiftl_t2n word_size) lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" apply (cases "n < size w") apply (erule neg_mask_is_div') apply (simp add: word_size) apply (frule p2_gt_0 [THEN Not_eq_iff [THEN iffD2], THEN iffD2]) apply (simp add: word_gt_0 del: p2_eq_0) apply (rule word_eqI) apply (simp add: word_ops_nth_size word_size) done lemma and_mask_arith': "0 < n \ w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" by (simp add: and_mask shiftr_div_2n_w shiftl_t2n word_size mult.commute) lemmas p2len = iffD2 [OF p2_eq_0 order_refl] lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" apply (cases "0 < n") apply (erule and_mask_arith') apply (simp add: word_size) apply (simp add: word_of_int_hom_syms word_div_def) done lemma mask_2pm1: "mask n = 2 ^ n - 1" by (simp add : mask_def) lemma add_mask_fold: "x + 2 ^ n - 1 = x + mask n" by (simp add: mask_def) lemma word_and_mask_le_2pm1: "w && mask n \ 2 ^ n - 1" by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u && mask n = 0 \ v < 2^n \ u && v = 0" apply (drule less_mask_eq) apply (simp add: mask_2pm1) apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (drule_tac x=na in word_eqD) apply (drule_tac x=na in word_eqD) apply simp done lemma le_shiftr1: "u <= v \ shiftr1 u <= shiftr1 v" apply (unfold word_le_def shiftr1_def word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (rename_tac bs bit) apply (case_tac "uint v" rule: bin_exhaust) apply (rename_tac bs' bit') apply (case_tac "bit") apply (case_tac "bit'", auto simp: less_eq_int_code le_Bits)[1] apply (case_tac bit') apply (simp add: le_Bits less_eq_int_code) apply (auto simp: le_Bits less_eq_int_code) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len0 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" apply (rule word_eqI) apply (simp add: word_size nth_shiftr) done lemmas shiftr_mask = order_refl [THEN shiftr_mask_le, simp] lemma word_leI: "(\n. \n < size (u::'a::len0 word); u !! n \ \ (v::'a::len0 word) !! n) \ u <= v" apply (rule xtr4) 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)" apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtr3) 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)" 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 (drule_tac f = "%u. u !! (x - n)" in arg_cong) apply (simp add : nth_shiftr) apply (case_tac "n <= x") apply auto done lemmas and_mask_eq_iff_le_mask = trans [OF and_mask_eq_iff_shiftr_0 le_mask_iff [THEN sym]] lemma mask_shiftl_decompose: "mask m << n = mask (m + n) && ~~ (mask n)" by (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) lemma one_bit_shiftl: "set_bit 0 n True = (1 :: 'a :: len word) << n" apply (rule word_eqI) apply (auto simp add: test_bit_set_gen nth_shiftl word_size simp del: word_set_bit_0 shiftl_1) done lemmas one_bit_pow = trans [OF one_bit_shiftl shiftl_1] lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] 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) = -(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 lemmas gt0_iff_gem1 = iffD1 [OF iffD1 [OF iff_left_commute le_m1_iff_lt] order_refl] lemmas power_2_ge_iff = trans [OF gt0_iff_gem1 [THEN sym] p2_gt_0] 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]) lemmas mask_lt_2pn = le_mask_iff_lt_2n [THEN iffD1, THEN iffD1, OF _ order_refl] lemma bang_eq: fixes x :: "'a::len0 word" shows "(x = y) = (\n. x !! n = y !! n)" by (subst test_bit_eq_iff[symmetric]) fastforce 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 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 word_and_max_word: fixes a::"'a::len word" shows "x = max_word \ a AND x = a" by simp -(* Simplifying with word_and_max_word and max_word_def works for arbitrary word sizes, - but the conditional rewrite can be slow when combined with other common rewrites on - word expressions. If we are willing to limit our attention to common word sizes, - the following will usually be much faster. *) -lemmas word_and_max_simps[simplified max_word_def, simplified] = - word_and_max[where 'a=8] word_and_max[where 'a=16] word_and_max[where 'a=32] word_and_max[where 'a=64] - lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [x !! 0]" by (simp add: word_and_1) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [x !! 0]" by (subst word_bw_comms) (simp add: word_and_1) lemma scast_scast_id [simp]: "scast (scast x :: ('a::len) signed word) = (x :: 'a word)" "scast (scast y :: ('a::len) word) = (y :: 'a signed word)" by (auto simp: is_up scast_up_scast_id) lemma scast_ucast_id [simp]: "scast (ucast (x :: 'a::len word) :: 'a signed word) = x" by (metis down_cast_same is_down len_signed order_refl scast_scast_id(1)) lemma ucast_scast_id [simp]: "ucast (scast (x :: 'a::len signed word) :: 'a word) = x" by (metis scast_scast_id(2) scast_ucast_id) lemma scast_of_nat [simp]: "scast (of_nat x :: 'a::len signed word) = (of_nat x :: 'a word)" by (metis (hide_lams, no_types) len_signed scast_def uint_sint word_of_nat word_ubin.Abs_norm word_ubin.eq_norm) lemma ucast_of_nat: "is_down (ucast :: 'a :: len word \ 'b :: len word) \ ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)" apply (rule sym) apply (subst word_unat.inverse_norm) apply (simp add: ucast_def word_of_int[symmetric] of_nat_nat[symmetric] unat_def[symmetric]) apply (simp add: unat_of_nat) apply (rule nat_int.Rep_eqD) apply (simp only: zmod_int) apply (rule mod_mod_cancel) by (simp add: is_down le_imp_power_dvd) (* shortcut for some specific lengths *) lemma word_fixed_sint_1[simp]: "sint (1::8 word) = 1" "sint (1::16 word) = 1" "sint (1::32 word) = 1" "sint (1::64 word) = 1" by (auto simp: sint_word_ariths) lemma word_sint_1 [simp]: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" apply (case_tac "LENGTH('a) \ 1") apply (metis diff_is_0_eq sbintrunc_0_numeral(1) sint_n1 word_1_wi word_m1_wi word_msb_1 word_msb_n1 word_sbin.Abs_norm) apply (metis bin_nth_1 diff_is_0_eq neq0_conv sbintrunc_minus_simps(4) sint_word_ariths(8) uint_1 word_msb_1 word_msb_nth) done lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (sbintrunc (LENGTH('a::len) - Suc 0) (1::int)))" by (metis One_nat_def scast_def sint_word_ariths(8)) lemma scast_1 [simp]: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (clarsimp simp: scast_1') (metis Suc_pred len_gt_0 nat.exhaust sbintrunc_Suc_numeral(1) uint_1 word_uint.Rep_inverse') lemma scast_eq_scast_id [simp]: "((scast (a :: 'a::len signed word) :: 'a word) = scast b) = (a = b)" by (metis ucast_scast_id) lemma ucast_eq_ucast_id [simp]: "((ucast (a :: 'a::len word) :: 'a signed word) = ucast b) = (a = b)" by (metis scast_ucast_id) lemma scast_ucast_norm [simp]: "(ucast (a :: 'a::len word) = (b :: 'a signed word)) = (a = scast b)" "((b :: 'a signed word) = ucast (a :: 'a::len word)) = (a = scast b)" by (metis scast_ucast_id ucast_scast_id)+ lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs && 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 of_int_uint [simp]: "of_int (uint x) = x" by (metis word_of_int word_uint.Rep_inverse') 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 corollary word_plus_and_or_coroll: "x && y = 0 \ x + y = x || y" using word_plus_and_or[where x=x and y=y] by simp corollary word_plus_and_or_coroll2: "(x && w) + (x && ~~ w) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply blast done lemma less_le_mult_nat': "w * c < b * c ==> 0 \ c ==> Suc w * c \ b * (c::nat)" apply (rule mult_right_mono) apply (rule Suc_leI) apply (erule (1) mult_right_less_imp_less) apply assumption done lemmas less_le_mult_nat = less_le_mult_nat'[simplified distrib_right, simplified] (* FIXME: these should eventually be moved to HOL/Word. *) lemmas extra_sle_sless_unfolds [simp] = word_sle_def[where a=0 and b=1] word_sle_def[where a=0 and b="numeral n"] word_sle_def[where a=1 and b=0] word_sle_def[where a=1 and b="numeral n"] word_sle_def[where a="numeral n" and b=0] word_sle_def[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 to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" proof - have "to_bl (1 :: 'a::len word) = to_bl (mask 1 :: 'a::len word)" by (simp add: mask_def) also have "\ = replicate (LENGTH('a) - 1) False @ [True]" by (cases "LENGTH('a)"; clarsimp simp: to_bl_mask) finally show ?thesis . qed lemma list_of_false: "True \ set xs \ xs = replicate (length xs) False" by (induct xs, simp_all) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" using list_of_false word_bl.Rep_inject by fastforce lemma diff_diff_less: "(i < m - (m - (n :: nat))) = (i < m \ i < n)" by auto lemma pop_count_0[simp]: "pop_count 0 = 0" by (clarsimp simp:pop_count_def) lemma pop_count_1[simp]: "pop_count 1 = 1" by (clarsimp simp:pop_count_def to_bl_1) lemma pop_count_0_imp_0: "(pop_count w = 0) = (w = 0)" apply (rule iffI) apply (clarsimp simp:pop_count_def) apply (subst (asm) filter_empty_conv) apply (clarsimp simp:eq_zero_set_bl) apply fast apply simp done end diff --git a/thys/Word_Lib/Word_Next.thy b/thys/Word_Lib/Word_Next.thy --- a/thys/Word_Lib/Word_Next.thy +++ b/thys/Word_Lib/Word_Next.thy @@ -1,86 +1,95 @@ (* SPDX-License-Identifier: BSD-3-Clause *) section\Increment and Decrement Machine Words Without Wrap-Around\ theory Word_Next imports Aligned begin text\Previous and next words addresses, without wrap around.\ definition word_next :: "'a::len word \ 'a::len word" where "word_next a \ if a = max_word then max_word else a + 1" definition word_prev :: "'a::len word \ 'a::len word" where "word_prev a \ if a = 0 then 0 else a - 1" text\Examples:\ lemma "word_next (2:: 8 word) = 3" by eval lemma "word_next (255:: 8 word) = 255" by eval lemma "word_prev (2:: 8 word) = 1" by eval lemma "word_prev (0:: 8 word) = 0" by eval 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 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)" 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" 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" by (meson not_less word_Suc_leq) -lemma word_lessThan_Suc_atMost: fixes k::"'a::len word" shows "k \ max_word \ {..< k + 1} = {..k}" - by(simp add: lessThan_def atMost_def 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: - fixes l::"'a::len word" shows "u \ max_word \ {l..< u + 1} = {l..u}" - by (simp add: atLeastAtMost_def atLeastLessThan_def word_lessThan_Suc_atMost) + \{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: fixes l::"'a::len word" - shows "m \ max_word \ {m<..u} = {m + 1..u}" - by(simp add: greaterThanAtMost_def greaterThan_def atLeastAtMost_def atLeast_def word_Suc_le) +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" 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::len0 word\ + by (auto simp add: le_less) + lemma word_adjacent_union: "word_next e = s' \ s \ e \ s' \ e' \ {s..e} \ {s'..e'} = {s .. e'}" - by (metis Un_absorb2 atLeastatMost_subset_iff ivl_disj_un_two(7) max_word_max - word_atLeastLessThan_Suc_atLeastAtMost word_le_less_eq word_next_def linorder_not_le) + apply (simp add: word_next_def ivl_disj_un_two_touch split: if_splits) + apply (drule sym) + apply simp + apply (subst word_atLeastLessThan_Suc_atLeastAtMost_union) + apply (simp_all add: word_Suc_le) + done end