lh-l4v/lib/Word_Lib/WordBitwise_Signed.thy

34 lines
932 B
Plaintext
Raw Normal View History

2016-04-17 19:33:19 +00:00
(*
* Copyright 2014, NICTA
*
* This software may be distributed and modified according to the terms of
* the BSD 2-Clause license. Note that NO WARRANTY is provided.
* See "LICENSE_BSD2.txt" for details.
*
* @TAG(NICTA_BSD)
*)
2016-05-13 12:41:20 +00:00
section "Bitwise tactic for Signed Words"
2016-04-17 19:33:19 +00:00
theory WordBitwise_Signed
imports
2019-05-19 07:52:18 +00:00
"HOL-Word.More_Word"
Signed_Words
begin
2016-05-13 11:23:02 +00:00
ML \<open>fun bw_tac_signed ctxt = let
val (ss, sss) = Word_Bitwise_Tac.expand_word_eq_sss
val sss = nth_map 2 (fn ss => put_simpset ss ctxt addsimps @{thms len_signed} |> simpset_of) sss
in
2018-06-09 09:41:48 +00:00
foldr1 (op THEN_ALL_NEW)
((CHANGED o safe_full_simp_tac (put_simpset ss ctxt)) ::
map (fn ss => safe_full_simp_tac (put_simpset ss ctxt)) sss)
end;
2016-05-13 11:23:02 +00:00
\<close>
method_setup word_bitwise_signed =
2016-05-13 11:23:02 +00:00
\<open>Scan.succeed (fn ctxt => Method.SIMPLE_METHOD (bw_tac_signed ctxt 1))\<close>
"decomposer for word equalities and inequalities into bit propositions"
end