diff --git a/lib/HaskellLib_H.thy b/lib/HaskellLib_H.thy index 59fcf4177..570839b15 100644 --- a/lib/HaskellLib_H.thy +++ b/lib/HaskellLib_H.thy @@ -12,7 +12,6 @@ theory HaskellLib_H imports Lib - NatBitwise More_Numeral_Type Monads.NonDetMonadVCG begin diff --git a/lib/Lib.thy b/lib/Lib.thy index 233de7e1c..487f8cf2a 100644 --- a/lib/Lib.thy +++ b/lib/Lib.thy @@ -2513,38 +2513,6 @@ lemma option_Some_value_independent: "\ f x = Some v; \v'. f x = Some v' \ f y = Some v' \ \ f y = Some v" by blast -text \Some int bitwise lemmas. Helpers for proofs about \<^file>\NatBitwise.thy\\ -lemma int_2p_eq_shiftl: - "(2::int)^x = 1 << x" - by (simp add: shiftl_int_def) - -lemma nat_int_mul: - "nat (int a * b) = a * nat b" - by (simp add: nat_mult_distrib) - -lemma int_shiftl_less_cancel: - "n \ m \ ((x :: int) << n < y << m) = (x < y << (m - n))" - apply (drule le_Suc_ex) - apply (clarsimp simp: shiftl_int_def power_add) - done - -lemma int_shiftl_lt_2p_bits: - "0 \ (x::int) \ x < 1 << n \ \i \ n. \ x !! i" - apply (clarsimp simp: shiftl_int_def) - by (metis bit_take_bit_iff not_less take_bit_int_eq_self_iff) -\ \TODO: The converse should be true as well, but seems hard to prove.\ - -lemmas int_eq_test_bit = bin_eq_iff -lemmas int_eq_test_bitI = int_eq_test_bit[THEN iffD2, rule_format] - -lemma le_nat_shrink_left: - "y \ z \ y = Suc x \ x < z" - by simp - -lemma length_ge_split: - "n < length xs \ \x xs'. xs = x # xs' \ n \ length xs'" - by (cases xs) auto - text \Support for defining enumerations on datatypes derived from enumerations\ lemma distinct_map_enum: "\ (\ x y. (F x = F y \ x = y )) \ diff --git a/lib/NatBitwise.thy b/tools/autocorres/NatBitwise.thy similarity index 62% rename from lib/NatBitwise.thy rename to tools/autocorres/NatBitwise.thy index c2ac5293c..7de19e00f 100644 --- a/lib/NatBitwise.thy +++ b/tools/autocorres/NatBitwise.thy @@ -4,11 +4,11 @@ * SPDX-License-Identifier: BSD-2-Clause *) -(* Instance of bit ops for nat. Used by HaskellLib and AutoCorres. - * Lemmas about this instance should also go here. *) +(* Instance of bit ops for nat. + Lemmas about this instance should also go here. *) theory NatBitwise imports - Lib + Word_Lib.WordSetup begin instantiation nat :: lsb @@ -50,10 +50,20 @@ lemma nat_2p_eq_shiftl: lemmas shiftl_nat_alt_def = shiftl_nat_def +lemma nat_int_mul: + "nat (int a * b) = a * nat b" + by (simp add: nat_mult_distrib) + lemma shiftl_nat_def: "(x::nat) << y = nat (int x << y)" by (simp add: nat_int_mul push_bit_eq_mult shiftl_def) +lemma int_shiftl_less_cancel: + "n \ m \ ((x :: int) << n < y << m) = (x < y << (m - n))" + apply (drule le_Suc_ex) + apply (clarsimp simp: shiftl_int_def power_add) + done + lemma nat_shiftl_less_cancel: "n \ m \ ((x :: nat) << n < y << m) = (x < y << (m - n))" apply (simp add: nat_int_comparison(2) shiftl_nat_def shiftl_def) @@ -69,4 +79,16 @@ lemma nat_shiftl_lt_2p_bits: lemmas nat_eq_test_bit = bit_eq_iff lemmas nat_eq_test_bitI = bit_eq_iff[THEN iffD2, rule_format] +lemma int_2p_eq_shiftl: + "(2::int)^x = 1 << x" + by (simp add: shiftl_int_def) + +lemma int_shiftl_lt_2p_bits: + "0 \ (x::int) \ x < 1 << n \ \i \ n. \ x !! i" + apply (clarsimp simp: shiftl_int_def) + by (metis bit_take_bit_iff not_less take_bit_int_eq_self_iff) +\ \TODO: The converse should be true as well, but seems hard to prove.\ + +lemmas int_eq_test_bitI = bin_eq_iff[THEN iffD2, rule_format] + end \ No newline at end of file diff --git a/tools/autocorres/WordAbstract.thy b/tools/autocorres/WordAbstract.thy index 5c58e4f84..d23fe0f36 100644 --- a/tools/autocorres/WordAbstract.thy +++ b/tools/autocorres/WordAbstract.thy @@ -8,7 +8,7 @@ theory WordAbstract imports L2Defs ExecConcrete - Lib.NatBitwise + NatBitwise begin definition "WORD_MAX x \ ((2 ^ (len_of x - 1) - 1) :: int)"