| author | haftmann | 
| Fri, 07 May 2010 15:05:52 +0200 | |
| changeset 36751 | 7f1da69cacb3 | 
| parent 33664 | d62805a237ef | 
| permissions | -rw-r--r-- | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 1 | (* Title: HOL/SMT/Tools/z3_model.ML | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 2 | Author: Sascha Boehme and Philipp Meyer, TU Muenchen | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 3 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 4 | Parser for counterexamples generated by Z3. | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 5 | *) | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 6 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 7 | signature Z3_MODEL = | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 8 | sig | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 9 | val parse_counterex: SMT_Translate.recon -> string list -> term list | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 10 | end | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 11 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 12 | structure Z3_Model: Z3_MODEL = | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 13 | struct | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 14 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 15 | (* counterexample expressions *) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 16 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 17 | datatype expr = True | False | Number of int * int option | Value of int | | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 18 | Array of array | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 19 | and array = Fresh of expr | Store of (array * expr) * expr | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 20 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 21 | |
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 22 | (* parsing *) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 23 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 24 | val space = Scan.many Symbol.is_ascii_blank | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 25 | fun in_parens p = Scan.$$ "(" |-- p --| Scan.$$ ")"
 | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 26 | fun in_braces p = (space -- Scan.$$ "{") |-- p --| (space -- Scan.$$ "}")
 | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 27 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 28 | val digit = (fn | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 29 | "0" => SOME 0 | "1" => SOME 1 | "2" => SOME 2 | "3" => SOME 3 | | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 30 | "4" => SOME 4 | "5" => SOME 5 | "6" => SOME 6 | "7" => SOME 7 | | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 31 | "8" => SOME 8 | "9" => SOME 9 | _ => NONE) | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 32 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 33 | val nat_num = Scan.repeat1 (Scan.some digit) >> | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 34 | (fn ds => fold (fn d => fn i => i * 10 + d) ds 0) | 
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 35 | val int_num = Scan.optional ($$ "-" >> K (fn i => ~i)) I :|-- | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 36 | (fn sign => nat_num >> sign) | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 37 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 38 | val is_char = Symbol.is_ascii_letter orf Symbol.is_ascii_digit orf | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 39 | member (op =) (explode "_+*-/%~=<>$&|?!.@^#") | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 40 | val name = Scan.many1 is_char >> implode | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 41 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 42 | fun array_expr st = st |> | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 43 | in_parens (space |-- ( | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 44 | Scan.this_string "const" |-- expr >> Fresh || | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 45 | Scan.this_string "store" -- space |-- array_expr -- expr -- expr >> Store)) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 46 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 47 | and expr st = st |> (space |-- ( | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 48 | Scan.this_string "true" >> K True || | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 49 | Scan.this_string "false" >> K False || | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 50 | int_num -- Scan.option (Scan.$$ "/" |-- int_num) >> Number || | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 51 | Scan.this_string "val!" |-- nat_num >> Value || | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 52 | array_expr >> Array)) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 53 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 54 | val mapping = space -- Scan.this_string "->" | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 55 | val value = mapping |-- expr | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 56 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 57 | val args_case = Scan.repeat expr -- value | 
| 33497 | 58 | val else_case = space -- Scan.this_string "else" |-- value >> | 
| 59 | pair ([] : expr list) | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 60 | |
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 61 | val func = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 62 | let fun cases st = (else_case >> single || args_case ::: cases) st | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 63 | in in_braces cases end | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 64 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 65 | val cex = space |-- Scan.repeat (space |-- name --| mapping -- | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 66 | (func || expr >> (single o pair []))) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 67 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 68 | fun read_cex ls = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 69 | explode (cat_lines ls) | 
| 33417 
ab2f6574a2a6
ignore parsing errors, return empty assignment instead
 boehmes parents: 
33378diff
changeset | 70 | |> try (fst o Scan.finite Symbol.stopper cex) | 
| 
ab2f6574a2a6
ignore parsing errors, return empty assignment instead
 boehmes parents: 
33378diff
changeset | 71 | |> the_default [] | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 72 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 73 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 74 | (* translation into terms *) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 75 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 76 | fun lookup_term tab (name, e) = Option.map (rpair e) (Symtab.lookup tab name) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 77 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 78 | fun with_name_context tab f xs = | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 79 | let | 
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 80 | val ns = Symtab.fold (Term.add_free_names o snd) tab [] | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 81 | val nctxt = Name.make_context ns | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 82 | in fst (fold_map f xs (Inttab.empty, nctxt)) end | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 83 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 84 | fun fresh_term T (tab, nctxt) = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 85 | let val (n, nctxt') = yield_singleton Name.variants "" nctxt | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 86 | in (Free (n, T), (tab, nctxt')) end | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 87 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 88 | fun term_of_value T i (cx as (tab, _)) = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 89 | (case Inttab.lookup tab i of | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 90 | SOME t => (t, cx) | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 91 | | NONE => | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 92 | let val (t, (tab', nctxt')) = fresh_term T cx | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 93 | in (t, (Inttab.update (i, t) tab', nctxt')) end) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 94 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 95 | fun trans_expr _ True = pair @{term True}
 | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 96 |   | trans_expr _ False = pair @{term False}
 | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 97 | | trans_expr T (Number (i, NONE)) = pair (HOLogic.mk_number T i) | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 98 | | trans_expr T (Number (i, SOME j)) = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 99 |       pair (Const (@{const_name divide}, [T, T] ---> T) $
 | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 100 | HOLogic.mk_number T i $ HOLogic.mk_number T j) | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 101 | | trans_expr T (Value i) = term_of_value T i | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 102 | | trans_expr T (Array a) = trans_array T a | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 103 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 104 | and trans_array T a = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 105 | let val dT = Term.domain_type T and rT = Term.range_type T | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 106 | in | 
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 107 | (case a of | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 108 |       Fresh e => trans_expr rT e #>> (fn t => Abs ("x", dT, t))
 | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 109 | | Store ((a', e1), e2) => | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 110 | trans_array T a' ##>> trans_expr dT e1 ##>> trans_expr rT e2 #>> | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 111 | (fn ((m, k), v) => | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 112 |           Const (@{const_name fun_upd}, [T, dT, rT] ---> T) $ m $ k $ v))
 | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 113 | end | 
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 114 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 115 | fun trans_pat i T f x = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 116 | f (Term.domain_type T) ##>> trans (i-1) (Term.range_type T) x #>> | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 117 | (fn (u, (us, t)) => (u :: us, t)) | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 118 | |
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 119 | and trans i T ([], v) = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 120 | if i > 0 then trans_pat i T fresh_term ([], v) | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 121 | else trans_expr T v #>> pair [] | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 122 | | trans i T (p :: ps, v) = trans_pat i T (fn U => trans_expr U p) (ps, v) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 123 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 124 | fun mk_eq' t us u = HOLogic.mk_eq (Term.list_comb (t, us), u) | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 125 | fun mk_eq (Const (@{const_name apply}, _)) (u' :: us', u) = mk_eq' u' us' u
 | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 126 | | mk_eq t (us, u) = mk_eq' t us u | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 127 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 128 | fun translate (t, cs) = | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 129 | let val T = Term.fastype_of t | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 130 | in | 
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 131 | (case (can HOLogic.dest_number t, cs) of | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 132 | (true, [c]) => trans 0 T c #>> (fn (_, u) => [mk_eq u ([], t)]) | 
| 33664 | 133 | | (_, (es, _) :: _) => fold_map (trans (length es) T) cs #>> map (mk_eq t) | 
| 134 |     | _ => raise TERM ("translate: no cases", [t]))
 | |
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 135 | end | 
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 136 | |
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 137 | |
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 138 | (* overall procedure *) | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 139 | |
| 33378 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 140 | fun parse_counterex ({terms, ...} : SMT_Translate.recon) ls =
 | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 141 | read_cex ls | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 142 | |> map_filter (lookup_term terms) | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 143 | |> with_name_context terms translate | 
| 
c394abc5f898
split parsing of counterexamples from translation into terms (avoids Term.dummyT and ill-typed terms)
 boehmes parents: 
33047diff
changeset | 144 | |> flat | 
| 32618 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 145 | |
| 
42865636d006
added new method "smt": an oracle-based connection to external SMT solvers
 boehmes parents: diff
changeset | 146 | end |