| author | wenzelm | 
| Wed, 19 Sep 2012 17:07:25 +0200 | |
| changeset 49445 | 638cefe3ee99 | 
| parent 46042 | ab32a87ba01a | 
| child 53998 | b352d3d4a58a | 
| permissions | -rw-r--r-- | 
| 40277 | 1 | (* Title: HOL/Tools/SMT/smt_builtin.ML | 
| 2 | Author: Sascha Boehme, TU Muenchen | |
| 3 | ||
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 4 | Tables of types and terms directly supported by SMT solvers. | 
| 40277 | 5 | *) | 
| 6 | ||
| 7 | signature SMT_BUILTIN = | |
| 8 | sig | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 9 | (*built-in types*) | 
| 41124 
1de17a2de5ad
moved SMT classes and dictionary functions to SMT_Utils
 boehmes parents: 
41073diff
changeset | 10 | val add_builtin_typ: SMT_Utils.class -> | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 11 | typ * (typ -> string option) * (typ -> int -> string option) -> | 
| 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 12 | Context.generic -> Context.generic | 
| 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 13 | val add_builtin_typ_ext: typ * (typ -> bool) -> Context.generic -> | 
| 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 14 | Context.generic | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 15 | val dest_builtin_typ: Proof.context -> typ -> string option | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 16 | val is_builtin_typ_ext: Proof.context -> typ -> bool | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 17 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 18 | (*built-in numbers*) | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 19 | val dest_builtin_num: Proof.context -> term -> (string * typ) option | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 20 | val is_builtin_num: Proof.context -> term -> bool | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 21 | val is_builtin_num_ext: Proof.context -> term -> bool | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 22 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 23 | (*built-in functions*) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 24 | type 'a bfun = Proof.context -> typ -> term list -> 'a | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 25 | type bfunr = string * int * term list * (term list -> term) | 
| 41124 
1de17a2de5ad
moved SMT classes and dictionary functions to SMT_Utils
 boehmes parents: 
41073diff
changeset | 26 | val add_builtin_fun: SMT_Utils.class -> | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 27 | (string * typ) * bfunr option bfun -> Context.generic -> Context.generic | 
| 41124 
1de17a2de5ad
moved SMT classes and dictionary functions to SMT_Utils
 boehmes parents: 
41073diff
changeset | 28 | val add_builtin_fun': SMT_Utils.class -> term * string -> Context.generic -> | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 29 | Context.generic | 
| 41354 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 30 | val add_builtin_fun_ext: (string * typ) * term list bfun -> | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 31 | Context.generic -> Context.generic | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 32 | val add_builtin_fun_ext': string * typ -> Context.generic -> Context.generic | 
| 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 33 | val add_builtin_fun_ext'': string -> Context.generic -> Context.generic | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 34 | val dest_builtin_fun: Proof.context -> string * typ -> term list -> | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 35 | bfunr option | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 36 | val dest_builtin_eq: Proof.context -> term -> term -> bfunr option | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 37 | val dest_builtin_pred: Proof.context -> string * typ -> term list -> | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 38 | bfunr option | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 39 | val dest_builtin_conn: Proof.context -> string * typ -> term list -> | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 40 | bfunr option | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 41 | val dest_builtin: Proof.context -> string * typ -> term list -> bfunr option | 
| 41354 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 42 | val dest_builtin_ext: Proof.context -> string * typ -> term list -> | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 43 | term list option | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 44 | val is_builtin_fun: Proof.context -> string * typ -> term list -> bool | 
| 41126 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 boehmes parents: 
41124diff
changeset | 45 | val is_builtin_fun_ext: Proof.context -> string * typ -> term list -> bool | 
| 40277 | 46 | end | 
| 47 | ||
| 48 | structure SMT_Builtin: SMT_BUILTIN = | |
| 49 | struct | |
| 50 | ||
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 51 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 52 | (* built-in tables *) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 53 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 54 | datatype ('a, 'b) kind = Ext of 'a | Int of 'b
 | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 55 | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
41281diff
changeset | 56 | type ('a, 'b) ttab = ((typ * ('a, 'b) kind) Ord_List.T) SMT_Utils.dict 
 | 
| 40677 
f5caf58e9cdb
more precise characterization of built-in constants "number_of", "0", and "1"
 blanchet parents: 
40664diff
changeset | 57 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 58 | fun typ_ord ((T, _), (U, _)) = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 59 | let | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 60 | fun tord (TVar _, Type _) = GREATER | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 61 | | tord (Type _, TVar _) = LESS | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 62 | | tord (Type (n, Ts), Type (m, Us)) = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 63 | if n = m then list_ord tord (Ts, Us) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 64 | else Term_Ord.typ_ord (T, U) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 65 | | tord TU = Term_Ord.typ_ord TU | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 66 | in tord (T, U) end | 
| 40277 | 67 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 68 | fun insert_ttab cs T f = | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
41281diff
changeset | 69 | SMT_Utils.dict_map_default (cs, []) | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 70 | (Ord_List.insert typ_ord (perhaps (try Logic.varifyT_global) T, f)) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 71 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 72 | fun merge_ttab ttabp = | 
| 41473 | 73 | SMT_Utils.dict_merge (Ord_List.merge typ_ord) ttabp | 
| 40277 | 74 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 75 | fun lookup_ttab ctxt ttab T = | 
| 42361 | 76 | let fun match (U, _) = Sign.typ_instance (Proof_Context.theory_of ctxt) (T, U) | 
| 41124 
1de17a2de5ad
moved SMT classes and dictionary functions to SMT_Utils
 boehmes parents: 
41073diff
changeset | 77 | in | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
41281diff
changeset | 78 | get_first (find_first match) | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
41281diff
changeset | 79 | (SMT_Utils.dict_lookup ttab (SMT_Config.solver_class_of ctxt)) | 
| 41124 
1de17a2de5ad
moved SMT classes and dictionary functions to SMT_Utils
 boehmes parents: 
41073diff
changeset | 80 | end | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 81 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 82 | type ('a, 'b) btab = ('a, 'b) ttab Symtab.table
 | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 83 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 84 | fun insert_btab cs n T f = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 85 | Symtab.map_default (n, []) (insert_ttab cs T f) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 86 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 87 | fun merge_btab btabp = Symtab.join (K merge_ttab) btabp | 
| 40277 | 88 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 89 | fun lookup_btab ctxt btab (n, T) = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 90 | (case Symtab.lookup btab n of | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 91 | NONE => NONE | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 92 | | SOME ttab => lookup_ttab ctxt ttab T) | 
| 40277 | 93 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 94 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 95 | (* built-in types *) | 
| 40277 | 96 | |
| 46042 | 97 | (* FIXME just one data slot (record) per program unit *) | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 98 | structure Builtin_Types = Generic_Data | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 99 | ( | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 100 | type T = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 101 | (typ -> bool, (typ -> string option) * (typ -> int -> string option)) ttab | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 102 | val empty = [] | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 103 | val extend = I | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 104 | val merge = merge_ttab | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 105 | ) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 106 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 107 | fun add_builtin_typ cs (T, f, g) = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 108 | Builtin_Types.map (insert_ttab cs T (Int (f, g))) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 109 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 110 | fun add_builtin_typ_ext (T, f) = | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
41281diff
changeset | 111 | Builtin_Types.map (insert_ttab SMT_Utils.basicC T (Ext f)) | 
| 40277 | 112 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 113 | fun lookup_builtin_typ ctxt = | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 114 | lookup_ttab ctxt (Builtin_Types.get (Context.Proof ctxt)) | 
| 40277 | 115 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 116 | fun dest_builtin_typ ctxt T = | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 117 | (case lookup_builtin_typ ctxt T of | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 118 | SOME (_, Int (f, _)) => f T | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 119 | | _ => NONE) | 
| 40277 | 120 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 121 | fun is_builtin_typ_ext ctxt T = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 122 | (case lookup_builtin_typ ctxt T of | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 123 | SOME (_, Int (f, _)) => is_some (f T) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 124 | | SOME (_, Ext f) => f T | 
| 40277 | 125 | | NONE => false) | 
| 126 | ||
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 127 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 128 | (* built-in numbers *) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 129 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 130 | fun dest_builtin_num ctxt t = | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 131 | (case try HOLogic.dest_number t of | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 132 | NONE => NONE | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 133 | | SOME (T, i) => | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 134 | (case lookup_builtin_typ ctxt T of | 
| 41127 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41126diff
changeset | 135 | SOME (_, Int (_, g)) => g T i |> Option.map (rpair T) | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 136 | | _ => NONE)) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 137 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 138 | val is_builtin_num = is_some oo dest_builtin_num | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 139 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 140 | fun is_builtin_num_ext ctxt t = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 141 | (case try HOLogic.dest_number t of | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 142 | NONE => false | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 143 | | SOME (T, _) => is_builtin_typ_ext ctxt T) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 144 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 145 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 146 | (* built-in functions *) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 147 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 148 | type 'a bfun = Proof.context -> typ -> term list -> 'a | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 149 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 150 | type bfunr = string * int * term list * (term list -> term) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 151 | |
| 46042 | 152 | (* FIXME just one data slot (record) per program unit *) | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 153 | structure Builtin_Funcs = Generic_Data | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 154 | ( | 
| 41354 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 155 | type T = (term list bfun, bfunr option bfun) btab | 
| 41126 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 boehmes parents: 
41124diff
changeset | 156 | val empty = Symtab.empty | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 157 | val extend = I | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 158 | val merge = merge_btab | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 159 | ) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 160 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 161 | fun add_builtin_fun cs ((n, T), f) = | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 162 | Builtin_Funcs.map (insert_btab cs n T (Int f)) | 
| 40277 | 163 | |
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 164 | fun add_builtin_fun' cs (t, n) = | 
| 41127 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 boehmes parents: 
41126diff
changeset | 165 | let | 
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 166 | val c as (m, T) = Term.dest_Const t | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 167 | fun app U ts = Term.list_comb (Const (m, U), ts) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 168 | fun bfun _ U ts = SOME (n, length (Term.binder_types T), ts, app U) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 169 | in add_builtin_fun cs (c, bfun) end | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 170 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 171 | fun add_builtin_fun_ext ((n, T), f) = | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
41281diff
changeset | 172 | Builtin_Funcs.map (insert_btab SMT_Utils.basicC n T (Ext f)) | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 173 | |
| 41354 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 174 | fun add_builtin_fun_ext' c = add_builtin_fun_ext (c, fn _ => fn _ => I) | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 175 | |
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 176 | fun add_builtin_fun_ext'' n context = | 
| 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 177 | let val thy = Context.theory_of context | 
| 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 178 | in add_builtin_fun_ext' (n, Sign.the_const_type thy n) context end | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 179 | |
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 180 | fun lookup_builtin_fun ctxt = | 
| 41072 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 boehmes parents: 
41059diff
changeset | 181 | lookup_btab ctxt (Builtin_Funcs.get (Context.Proof ctxt)) | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 182 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 183 | fun dest_builtin_fun ctxt (c as (_, T)) ts = | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 184 | (case lookup_builtin_fun ctxt c of | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 185 | SOME (_, Int f) => f ctxt T ts | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 186 | | _ => NONE) | 
| 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 187 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 188 | fun dest_builtin_eq ctxt t u = | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 189 | let | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 190 |     val aT = TFree (Name.aT, @{sort type})
 | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 191 |     val c = (@{const_name HOL.eq}, aT --> aT --> @{typ bool})
 | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 192 | fun mk ts = Term.list_comb (HOLogic.eq_const (Term.fastype_of (hd ts)), ts) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 193 | in | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 194 | dest_builtin_fun ctxt c [] | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 195 | |> Option.map (fn (n, i, _, _) => (n, i, [t, u], mk)) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 196 | end | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 197 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 198 | fun special_builtin_fun pred ctxt (c as (_, 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: 
41280diff
changeset | 199 | if pred (Term.body_type T, Term.binder_types T) then | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 200 | dest_builtin_fun ctxt c ts | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 201 | 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: 
41280diff
changeset | 202 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 203 | fun dest_builtin_pred ctxt = special_builtin_fun (equal @{typ bool} o fst) ctxt
 | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 204 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 205 | fun dest_builtin_conn ctxt = | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 206 |   special_builtin_fun (forall (equal @{typ bool}) o (op ::)) ctxt
 | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 207 | |
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 208 | fun dest_builtin ctxt c ts = | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 209 | let val t =Term.list_comb (Const c, ts) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 210 | in | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 211 | (case dest_builtin_num ctxt t of | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 212 | SOME (n, _) => SOME (n, 0, [], K t) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 213 | | NONE => dest_builtin_fun ctxt c ts) | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 214 | end | 
| 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 215 | |
| 41354 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 216 | fun dest_builtin_fun_ext ctxt (c as (_, T)) ts = | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 217 | (case lookup_builtin_fun ctxt c of | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 218 | SOME (_, Int f) => f ctxt T ts |> Option.map (fn (_, _, us, _) => us) | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 219 | | SOME (_, Ext f) => SOME (f ctxt T ts) | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 220 | | NONE => NONE) | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 221 | |
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 222 | fun dest_builtin_ext ctxt c ts = | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 223 | if is_builtin_num_ext ctxt (Term.list_comb (Const c, ts)) then SOME [] | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 224 | else dest_builtin_fun_ext ctxt c ts | 
| 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 225 | |
| 41281 
679118e35378
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
 boehmes parents: 
41280diff
changeset | 226 | fun is_builtin_fun ctxt c ts = is_some (dest_builtin_fun ctxt c ts) | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 227 | |
| 41354 
0abe5db19f3a
also provide a view on arguments for "external" built-in symbols (similar to "internal" (real) built-in symbols)
 boehmes parents: 
41336diff
changeset | 228 | fun is_builtin_fun_ext ctxt c ts = is_some (dest_builtin_fun_ext ctxt c ts) | 
| 41059 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 boehmes parents: 
40691diff
changeset | 229 | |
| 40277 | 230 | end |