blanchet@58058: (* Title: HOL/Library/Old_SMT/old_smt_word.ML boehmes@41060: Author: Sascha Boehme, TU Muenchen boehmes@41060: boehmes@41060: SMT setup for words. boehmes@41060: *) boehmes@41060: blanchet@58058: signature OLD_SMT_WORD = boehmes@41060: sig boehmes@41060: val setup: theory -> theory boehmes@41060: end boehmes@41060: blanchet@58058: structure Old_SMT_Word: OLD_SMT_WORD = boehmes@41060: struct boehmes@41060: thomas@47567: open Word_Lib boehmes@41060: boehmes@41060: (* SMT-LIB logic *) boehmes@41060: boehmes@41060: fun smtlib_logic ts = boehmes@41060: if exists (Term.exists_type (Term.exists_subtype is_wordT)) ts boehmes@41060: then SOME "QF_AUFBV" boehmes@41060: else NONE boehmes@41060: boehmes@41060: boehmes@41060: (* SMT-LIB builtins *) boehmes@41060: boehmes@41060: local blanchet@58058: val smtlibC = Old_SMTLIB_Interface.smtlibC boehmes@41061: boehmes@41061: val wordT = @{typ "'a::len word"} boehmes@41061: boehmes@41060: fun index1 n i = n ^ "[" ^ string_of_int i ^ "]" boehmes@41060: fun index2 n i j = n ^ "[" ^ string_of_int i ^ ":" ^ string_of_int j ^ "]" boehmes@41060: boehmes@41061: fun word_typ (Type (@{type_name word}, [T])) = boehmes@41060: Option.map (index1 "BitVec") (try dest_binT T) boehmes@41061: | word_typ _ = NONE boehmes@41060: boehmes@41061: fun word_num (Type (@{type_name word}, [T])) i = boehmes@41060: Option.map (index1 ("bv" ^ string_of_int i)) (try dest_binT T) boehmes@41061: | word_num _ _ = NONE boehmes@41060: boehmes@41281: fun if_fixed pred m n T ts = boehmes@41127: let val (Us, U) = Term.strip_type T boehmes@41127: in boehmes@41281: if pred (U, Us) then boehmes@41281: SOME (n, length Us, ts, Term.list_comb o pair (Const (m, T))) boehmes@41127: else NONE boehmes@41127: end boehmes@41061: boehmes@41281: fun if_fixed_all m = if_fixed (forall (can dest_wordT) o (op ::)) m boehmes@41281: fun if_fixed_args m = if_fixed (forall (can dest_wordT) o snd) m boehmes@41281: boehmes@41061: fun add_word_fun f (t, n) = wenzelm@46124: let val (m, _) = Term.dest_Const t blanchet@58058: in Old_SMT_Builtin.add_builtin_fun smtlibC (Term.dest_Const t, K (f m n)) end boehmes@41281: boehmes@41281: fun hd2 xs = hd (tl xs) boehmes@41061: boehmes@41281: fun mk_nat i = @{const nat} $ HOLogic.mk_number @{typ nat} i boehmes@41281: boehmes@41281: fun dest_nat (@{const nat} $ n) = snd (HOLogic.dest_number n) boehmes@41281: | dest_nat t = raise TERM ("not a natural number", [t]) boehmes@41281: boehmes@41281: fun mk_shift c [t, u] = Const c $ t $ mk_nat (snd (HOLogic.dest_number u)) boehmes@41281: | mk_shift c ts = raise TERM ("bad arguments", Const c :: ts) boehmes@41061: boehmes@41281: fun shift m n T ts = boehmes@41281: let val U = Term.domain_type T boehmes@41281: in boehmes@41281: (case (can dest_wordT U, try (dest_nat o hd2) ts) of boehmes@41281: (true, SOME i) => boehmes@41281: SOME (n, 2, [hd ts, HOLogic.mk_number U i], mk_shift (m, T)) boehmes@41281: | _ => NONE) (* FIXME: also support non-numerical shifts *) boehmes@41281: end boehmes@41060: boehmes@41281: fun mk_extract c i ts = Term.list_comb (Const c, mk_nat i :: ts) boehmes@41281: boehmes@41281: fun extract m n T ts = boehmes@41281: let val U = Term.range_type (Term.range_type T) boehmes@41060: in boehmes@41281: (case (try (dest_nat o hd) ts, try dest_wordT U) of boehmes@41281: (SOME lb, SOME i) => boehmes@41281: SOME (index2 n (i + lb - 1) lb, 1, tl ts, mk_extract (m, T) lb) boehmes@41060: | _ => NONE) boehmes@41060: end boehmes@41060: boehmes@41281: fun mk_extend c ts = Term.list_comb (Const c, ts) boehmes@41061: boehmes@41281: fun extend m n T ts = boehmes@41281: let val (U1, U2) = Term.dest_funT T boehmes@41281: in boehmes@41281: (case (try dest_wordT U1, try dest_wordT U2) of boehmes@41281: (SOME i, SOME j) => boehmes@41281: if j-i >= 0 then SOME (index1 n (j-i), 1, ts, mk_extend (m, T)) boehmes@41281: else NONE boehmes@41281: | _ => NONE) boehmes@41281: end boehmes@41060: boehmes@41281: fun mk_rotate c i ts = Term.list_comb (Const c, mk_nat i :: ts) boehmes@41281: boehmes@41281: fun rotate m n T ts = boehmes@41281: let val U = Term.domain_type (Term.range_type T) boehmes@41127: in boehmes@41281: (case (can dest_wordT U, try (dest_nat o hd) ts) of boehmes@41281: (true, SOME i) => SOME (index1 n i, 1, tl ts, mk_rotate (m, T) i) boehmes@41281: | _ => NONE) boehmes@41127: end boehmes@41060: in boehmes@41060: boehmes@41061: val setup_builtins = blanchet@58058: Old_SMT_Builtin.add_builtin_typ smtlibC (wordT, word_typ, word_num) #> boehmes@41281: fold (add_word_fun if_fixed_all) [ boehmes@41061: (@{term "uminus :: 'a::len word => _"}, "bvneg"), boehmes@41061: (@{term "plus :: 'a::len word => _"}, "bvadd"), boehmes@41061: (@{term "minus :: 'a::len word => _"}, "bvsub"), boehmes@41061: (@{term "times :: 'a::len word => _"}, "bvmul"), boehmes@41061: (@{term "bitNOT :: 'a::len word => _"}, "bvnot"), boehmes@41061: (@{term "bitAND :: 'a::len word => _"}, "bvand"), boehmes@41061: (@{term "bitOR :: 'a::len word => _"}, "bvor"), boehmes@41061: (@{term "bitXOR :: 'a::len word => _"}, "bvxor"), boehmes@41061: (@{term "word_cat :: 'a::len word => _"}, "concat") ] #> boehmes@41061: fold (add_word_fun shift) [ boehmes@41061: (@{term "shiftl :: 'a::len word => _ "}, "bvshl"), boehmes@41061: (@{term "shiftr :: 'a::len word => _"}, "bvlshr"), boehmes@41061: (@{term "sshiftr :: 'a::len word => _"}, "bvashr") ] #> boehmes@41061: add_word_fun extract boehmes@41061: (@{term "slice :: _ => 'a::len word => _"}, "extract") #> boehmes@41061: fold (add_word_fun extend) [ boehmes@41061: (@{term "ucast :: 'a::len word => _"}, "zero_extend"), boehmes@41061: (@{term "scast :: 'a::len word => _"}, "sign_extend") ] #> boehmes@41061: fold (add_word_fun rotate) [ boehmes@41061: (@{term word_rotl}, "rotate_left"), boehmes@41061: (@{term word_rotr}, "rotate_right") ] #> boehmes@41281: fold (add_word_fun if_fixed_args) [ boehmes@41061: (@{term "less :: 'a::len word => _"}, "bvult"), boehmes@41061: (@{term "less_eq :: 'a::len word => _"}, "bvule"), boehmes@41061: (@{term word_sless}, "bvslt"), boehmes@41061: (@{term word_sle}, "bvsle") ] boehmes@41060: boehmes@41060: end boehmes@41060: boehmes@41060: boehmes@41060: (* setup *) boehmes@41060: boehmes@41060: val setup = boehmes@41072: Context.theory_map ( blanchet@58058: Old_SMTLIB_Interface.add_logic (20, smtlib_logic) #> boehmes@41072: setup_builtins) boehmes@41060: boehmes@41060: end