| author | wenzelm | 
| Sun, 18 Dec 2016 13:46:57 +0100 | |
| changeset 64597 | 1c252d8b6ca6 | 
| parent 58825 | 2065f49da190 | 
| permissions | -rw-r--r-- | 
| 58058 
1a0b18176548
add 'old_' prefix to SMT file names; add 'Old_' to ML module names;
 blanchet parents: 
58057diff
changeset | 1 | (* Title: HOL/Library/Old_SMT/old_smt_word.ML | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 2 | Author: Sascha Boehme, TU Muenchen | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 3 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 4 | SMT setup for words. | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 5 | *) | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 6 | |
| 58825 | 7 | structure Old_SMT_Word: sig end = | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 8 | struct | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 9 | |
| 47567 
407cabf66f21
New tactic "word_bitwise" expands word equalities/inequalities into logic.
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
46124diff
changeset | 10 | open Word_Lib | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 11 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 12 | (* SMT-LIB logic *) | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 13 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 14 | fun smtlib_logic ts = | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 15 | if exists (Term.exists_type (Term.exists_subtype is_wordT)) ts | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 16 | then SOME "QF_AUFBV" | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 17 | else NONE | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 18 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 19 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 20 | (* SMT-LIB builtins *) | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 21 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 22 | local | 
| 58058 
1a0b18176548
add 'old_' prefix to SMT file names; add 'Old_' to ML module names;
 blanchet parents: 
58057diff
changeset | 23 | val smtlibC = Old_SMTLIB_Interface.smtlibC | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 24 | |
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 25 |   val wordT = @{typ "'a::len word"}
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 26 | |
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 27 | fun index1 n i = n ^ "[" ^ string_of_int i ^ "]" | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 28 | fun index2 n i j = n ^ "[" ^ string_of_int i ^ ":" ^ string_of_int j ^ "]" | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 29 | |
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 30 |   fun word_typ (Type (@{type_name word}, [T])) =
 | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 31 | Option.map (index1 "BitVec") (try dest_binT T) | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 32 | | word_typ _ = NONE | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 33 | |
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 34 |   fun word_num (Type (@{type_name word}, [T])) i =
 | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 35 |         Option.map (index1 ("bv" ^ string_of_int i)) (try dest_binT T)
 | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 36 | | word_num _ _ = NONE | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 37 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 38 | fun if_fixed pred m n T ts = | 
| 41127 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41072diff
changeset | 39 | let val (Us, U) = Term.strip_type T | 
| 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41072diff
changeset | 40 | in | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 41 | if pred (U, Us) then | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 42 | SOME (n, length Us, ts, Term.list_comb o pair (Const (m, T))) | 
| 41127 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41072diff
changeset | 43 | else NONE | 
| 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41072diff
changeset | 44 | end | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 45 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 46 | fun if_fixed_all m = if_fixed (forall (can dest_wordT) o (op ::)) m | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 47 | fun if_fixed_args m = if_fixed (forall (can dest_wordT) o snd) m | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 48 | |
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 49 | fun add_word_fun f (t, n) = | 
| 46124 | 50 | let val (m, _) = Term.dest_Const t | 
| 58058 
1a0b18176548
add 'old_' prefix to SMT file names; add 'Old_' to ML module names;
 blanchet parents: 
58057diff
changeset | 51 | in Old_SMT_Builtin.add_builtin_fun smtlibC (Term.dest_Const t, K (f m n)) end | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 52 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 53 | fun hd2 xs = hd (tl xs) | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 54 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 55 |   fun mk_nat i = @{const nat} $ HOLogic.mk_number @{typ nat} i
 | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 56 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 57 |   fun dest_nat (@{const nat} $ n) = snd (HOLogic.dest_number n)
 | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 58 |     | dest_nat t = raise TERM ("not a natural number", [t])
 | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 59 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 60 | fun mk_shift c [t, u] = Const c $ t $ mk_nat (snd (HOLogic.dest_number u)) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 61 |     | mk_shift c ts = raise TERM ("bad arguments", Const c :: ts)
 | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 62 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 63 | fun shift m n T ts = | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 64 | let val U = Term.domain_type T | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 65 | in | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 66 | (case (can dest_wordT U, try (dest_nat o hd2) ts) of | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 67 | (true, SOME i) => | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 68 | SOME (n, 2, [hd ts, HOLogic.mk_number U i], mk_shift (m, T)) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 69 | | _ => NONE) (* FIXME: also support non-numerical shifts *) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 70 | end | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 71 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 72 | fun mk_extract c i ts = Term.list_comb (Const c, mk_nat i :: ts) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 73 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 74 | fun extract m n T ts = | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 75 | let val U = Term.range_type (Term.range_type T) | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 76 | in | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 77 | (case (try (dest_nat o hd) ts, try dest_wordT U) of | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 78 | (SOME lb, SOME i) => | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 79 | SOME (index2 n (i + lb - 1) lb, 1, tl ts, mk_extract (m, T) lb) | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 80 | | _ => NONE) | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 81 | end | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 82 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 83 | fun mk_extend c ts = Term.list_comb (Const c, ts) | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 84 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 85 | fun extend m n T ts = | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 86 | let val (U1, U2) = Term.dest_funT T | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 87 | in | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 88 | (case (try dest_wordT U1, try dest_wordT U2) of | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 89 | (SOME i, SOME j) => | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 90 | if j-i >= 0 then SOME (index1 n (j-i), 1, ts, mk_extend (m, T)) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 91 | else NONE | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 92 | | _ => NONE) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 93 | end | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 94 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 95 | fun mk_rotate c i ts = Term.list_comb (Const c, mk_nat i :: ts) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 96 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 97 | fun rotate m n T ts = | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 98 | let val U = Term.domain_type (Term.range_type T) | 
| 41127 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41072diff
changeset | 99 | in | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 100 | (case (can dest_wordT U, try (dest_nat o hd) ts) of | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 101 | (true, SOME i) => SOME (index1 n i, 1, tl ts, mk_rotate (m, T) i) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 102 | | _ => NONE) | 
| 41127 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41072diff
changeset | 103 | end | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 104 | in | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 105 | |
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 106 | val setup_builtins = | 
| 58058 
1a0b18176548
add 'old_' prefix to SMT file names; add 'Old_' to ML module names;
 blanchet parents: 
58057diff
changeset | 107 | Old_SMT_Builtin.add_builtin_typ smtlibC (wordT, word_typ, word_num) #> | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 108 | fold (add_word_fun if_fixed_all) [ | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 109 |     (@{term "uminus :: 'a::len word => _"}, "bvneg"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 110 |     (@{term "plus :: 'a::len word => _"}, "bvadd"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 111 |     (@{term "minus :: 'a::len word => _"}, "bvsub"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 112 |     (@{term "times :: 'a::len word => _"}, "bvmul"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 113 |     (@{term "bitNOT :: 'a::len word => _"}, "bvnot"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 114 |     (@{term "bitAND :: 'a::len word => _"}, "bvand"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 115 |     (@{term "bitOR :: 'a::len word => _"}, "bvor"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 116 |     (@{term "bitXOR :: 'a::len word => _"}, "bvxor"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 117 |     (@{term "word_cat :: 'a::len word => _"}, "concat") ] #>
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 118 | fold (add_word_fun shift) [ | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 119 |     (@{term "shiftl :: 'a::len word => _ "}, "bvshl"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 120 |     (@{term "shiftr :: 'a::len word => _"}, "bvlshr"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 121 |     (@{term "sshiftr :: 'a::len word => _"}, "bvashr") ] #>
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 122 | add_word_fun extract | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 123 |     (@{term "slice :: _ => 'a::len word => _"}, "extract") #>
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 124 | fold (add_word_fun extend) [ | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 125 |     (@{term "ucast :: 'a::len word => _"}, "zero_extend"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 126 |     (@{term "scast :: 'a::len word => _"}, "sign_extend") ] #>
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 127 | fold (add_word_fun rotate) [ | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 128 |     (@{term word_rotl}, "rotate_left"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 129 |     (@{term word_rotr}, "rotate_right") ] #>
 | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41127diff
changeset | 130 | fold (add_word_fun if_fixed_args) [ | 
| 41061 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 131 |     (@{term "less :: 'a::len word => _"}, "bvult"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 132 |     (@{term "less_eq :: 'a::len word => _"}, "bvule"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 133 |     (@{term word_sless}, "bvslt"),
 | 
| 
492f8fd35fc0
centralized handling of built-in types and constants for bitvectors
 boehmes parents: 
41060diff
changeset | 134 |     (@{term word_sle}, "bvsle") ]
 | 
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 135 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 136 | end | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 137 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 138 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 139 | (* setup *) | 
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 140 | |
| 58825 | 141 | val _ = | 
| 142 | Theory.setup | |
| 143 | (Context.theory_map | |
| 144 | (Old_SMTLIB_Interface.add_logic (20, smtlib_logic) #> setup_builtins)) | |
| 41060 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 145 | |
| 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 boehmes parents: diff
changeset | 146 | end |