| author | boehmes | 
| Sun, 19 Dec 2010 17:55:56 +0100 | |
| changeset 41280 | a7de9d36f4f2 | 
| parent 41127 | 2ea84c8535c6 | 
| child 41302 | 0485186839a7 | 
| permissions | -rw-r--r-- | 
| 36898 | 1  | 
(* Title: HOL/Tools/SMT/z3_interface.ML  | 
2  | 
Author: Sascha Boehme, TU Muenchen  | 
|
3  | 
||
4  | 
Interface to Z3 based on a relaxed version of SMT-LIB.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature Z3_INTERFACE =  | 
|
8  | 
sig  | 
|
| 
41124
 
1de17a2de5ad
moved SMT classes and dictionary functions to SMT_Utils
 
boehmes 
parents: 
41072 
diff
changeset
 | 
9  | 
val smtlib_z3C: SMT_Utils.class  | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
10  | 
val setup: theory -> theory  | 
| 36898 | 11  | 
|
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
12  | 
datatype sym = Sym of string * sym list  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
13  | 
  type mk_builtins = {
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
14  | 
mk_builtin_typ: sym -> typ option,  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
15  | 
mk_builtin_num: theory -> int -> typ -> cterm option,  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
16  | 
mk_builtin_fun: theory -> sym -> cterm list -> cterm option }  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
17  | 
val add_mk_builtins: mk_builtins -> Context.generic -> Context.generic  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
18  | 
val mk_builtin_typ: Proof.context -> sym -> typ option  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
19  | 
val mk_builtin_num: Proof.context -> int -> typ -> cterm option  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
20  | 
val mk_builtin_fun: Proof.context -> sym -> cterm list -> cterm option  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
21  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
22  | 
val is_builtin_theory_term: Proof.context -> term -> bool  | 
| 36898 | 23  | 
end  | 
24  | 
||
25  | 
structure Z3_Interface: Z3_INTERFACE =  | 
|
26  | 
struct  | 
|
27  | 
||
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
28  | 
structure U = SMT_Utils  | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
29  | 
structure B = SMT_Builtin  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
30  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
31  | 
val smtlib_z3C = SMTLIB_Interface.smtlibC @ ["z3"]  | 
| 36898 | 32  | 
|
33  | 
||
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
34  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
35  | 
(* interface *)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
36  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
37  | 
local  | 
| 
41127
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
38  | 
fun translate_config ctxt =  | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
39  | 
let  | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
40  | 
      val {prefixes, header, is_fol, serialize, ...} =
 | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
41  | 
SMTLIB_Interface.translate_config ctxt  | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
42  | 
in  | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
43  | 
      {prefixes=prefixes, header=header, is_fol=is_fol, serialize=serialize,
 | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
44  | 
has_datatypes=true}  | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
45  | 
end  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
46  | 
|
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
47  | 
  fun is_div_mod (@{const div (int)} $ _ $ t) = U.is_number t
 | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
48  | 
    | is_div_mod (@{const mod (int)} $ _ $ t) = U.is_number t
 | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
49  | 
| is_div_mod _ = false  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
50  | 
|
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
51  | 
  val div_by_z3div = mk_meta_eq @{lemma
 | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
52  | 
"k div l = (  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
53  | 
if k = 0 | l = 0 then 0  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
54  | 
else if (0 < k & 0 < l) | (k < 0 & 0 < l) then z3div k l  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
55  | 
else z3div (-k) (-l))"  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
56  | 
by (simp add: SMT.z3div_def)}  | 
| 37151 | 57  | 
|
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
58  | 
  val mod_by_z3mod = mk_meta_eq @{lemma
 | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
59  | 
"k mod l = (  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
60  | 
if l = 0 then k  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
61  | 
else if k = 0 then 0  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
62  | 
else if (0 < k & 0 < l) | (k < 0 & 0 < l) then z3mod k l  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
63  | 
else - z3mod (-k) (-l))"  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
64  | 
by (simp add: z3mod_def)}  | 
| 37151 | 65  | 
|
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
66  | 
fun div_mod_conv _ =  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
67  | 
U.if_true_conv is_div_mod (Conv.rewrs_conv [div_by_z3div, mod_by_z3mod])  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
68  | 
|
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
69  | 
fun rewrite_div_mod ctxt thm =  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
70  | 
if Term.exists_subterm is_div_mod (Thm.prop_of thm) then  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
71  | 
Conv.fconv_rule (Conv.top_conv div_mod_conv ctxt) thm  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
72  | 
else thm  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
73  | 
|
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
74  | 
fun norm_div_mod ctxt = pairself (map (rewrite_div_mod ctxt))  | 
| 
41072
 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 
boehmes 
parents: 
41059 
diff
changeset
 | 
75  | 
|
| 
 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 
boehmes 
parents: 
41059 
diff
changeset
 | 
76  | 
val setup_builtins =  | 
| 
 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 
boehmes 
parents: 
41059 
diff
changeset
 | 
77  | 
    B.add_builtin_fun' smtlib_z3C (@{const z3div}, "div") #>
 | 
| 
 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 
boehmes 
parents: 
41059 
diff
changeset
 | 
78  | 
    B.add_builtin_fun' smtlib_z3C (@{const z3mod}, "mod")
 | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
79  | 
in  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
80  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41124 
diff
changeset
 | 
81  | 
val setup = Context.theory_map (  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41124 
diff
changeset
 | 
82  | 
setup_builtins #>  | 
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41127 
diff
changeset
 | 
83  | 
SMT_Normalize.add_extra_norm (smtlib_z3C, norm_div_mod) #>  | 
| 
41127
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
84  | 
SMT_Translate.add_config (smtlib_z3C, translate_config))  | 
| 
41072
 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 
boehmes 
parents: 
41059 
diff
changeset
 | 
85  | 
|
| 36898 | 86  | 
end  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
87  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
88  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
89  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
90  | 
(* constructors *)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
91  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
92  | 
datatype sym = Sym of string * sym list  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
93  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
94  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
95  | 
(** additional constructors **)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
96  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
97  | 
type mk_builtins = {
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
98  | 
mk_builtin_typ: sym -> typ option,  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
99  | 
mk_builtin_num: theory -> int -> typ -> cterm option,  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
100  | 
mk_builtin_fun: theory -> sym -> cterm list -> cterm option }  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
101  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
102  | 
fun chained _ [] = NONE  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
103  | 
| chained f (b :: bs) = (case f b of SOME y => SOME y | NONE => chained f bs)  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
104  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
105  | 
fun chained_mk_builtin_typ bs sym =  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
106  | 
  chained (fn {mk_builtin_typ=mk, ...} : mk_builtins => mk sym) bs
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
107  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
108  | 
fun chained_mk_builtin_num ctxt bs i T =  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
109  | 
let val thy = ProofContext.theory_of ctxt  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
110  | 
  in chained (fn {mk_builtin_num=mk, ...} : mk_builtins => mk thy i T) bs end
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
111  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
112  | 
fun chained_mk_builtin_fun ctxt bs s cts =  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
113  | 
let val thy = ProofContext.theory_of ctxt  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
114  | 
  in chained (fn {mk_builtin_fun=mk, ...} : mk_builtins => mk thy s cts) bs end
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
115  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
116  | 
fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2)  | 
| 
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
117  | 
|
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
118  | 
structure Mk_Builtins = Generic_Data  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
119  | 
(  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
120  | 
type T = (int * mk_builtins) list  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
121  | 
val empty = []  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
122  | 
val extend = I  | 
| 39687 | 123  | 
fun merge (bs1, bs2) = Ord_List.union fst_int_ord bs2 bs1  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
124  | 
)  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
125  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
126  | 
fun add_mk_builtins mk =  | 
| 39687 | 127  | 
Mk_Builtins.map (Ord_List.insert fst_int_ord (serial (), mk))  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
128  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
129  | 
fun get_mk_builtins ctxt = map snd (Mk_Builtins.get (Context.Proof ctxt))  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
130  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
131  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
132  | 
(** basic and additional constructors **)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
133  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
134  | 
fun mk_builtin_typ _ (Sym ("bool", _)) = SOME @{typ bool}
 | 
| 40516 | 135  | 
  | mk_builtin_typ _ (Sym ("Int", _)) = SOME @{typ int}
 | 
136  | 
  | mk_builtin_typ _ (Sym ("int", _)) = SOME @{typ int}  (*FIXME: delete*)
 | 
|
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
137  | 
| mk_builtin_typ ctxt sym = chained_mk_builtin_typ (get_mk_builtins ctxt) sym  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
138  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
139  | 
fun mk_builtin_num _ i @{typ int} = SOME (Numeral.mk_cnumber @{ctyp int} i)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
140  | 
| mk_builtin_num ctxt i T =  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
141  | 
chained_mk_builtin_num ctxt (get_mk_builtins ctxt) i T  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
142  | 
|
| 
40579
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
143  | 
val mk_true = Thm.cterm_of @{theory} (@{const Not} $ @{const False})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
144  | 
val mk_false = Thm.cterm_of @{theory} @{const False}
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
145  | 
val mk_not = Thm.capply (Thm.cterm_of @{theory} @{const Not})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
146  | 
val mk_implies = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.implies})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
147  | 
val mk_iff = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.eq (bool)})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
148  | 
val conj = Thm.cterm_of @{theory} @{const HOL.conj}
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
149  | 
val disj = Thm.cterm_of @{theory} @{const HOL.disj}
 | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
150  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
151  | 
fun mk_nary _ cu [] = cu  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
152  | 
| mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts)  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
153  | 
|
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
154  | 
val eq = U.mk_const_pat @{theory} @{const_name HOL.eq} U.destT1
 | 
| 
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
155  | 
fun mk_eq ct cu = Thm.mk_binop (U.instT' ct eq) ct cu  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
156  | 
|
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
157  | 
val if_term = U.mk_const_pat @{theory} @{const_name If} (U.destT1 o U.destT2)
 | 
| 
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
158  | 
fun mk_if cc ct cu = Thm.mk_binop (Thm.capply (U.instT' ct if_term) cc) ct cu  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
159  | 
|
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
160  | 
val nil_term = U.mk_const_pat @{theory} @{const_name Nil} U.destT1
 | 
| 
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
161  | 
val cons_term = U.mk_const_pat @{theory} @{const_name Cons} U.destT1
 | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
162  | 
fun mk_list cT cts =  | 
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
163  | 
fold_rev (Thm.mk_binop (U.instT cT cons_term)) cts (U.instT cT nil_term)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
164  | 
|
| 
40681
 
872b08416fb4
be more precise: only treat constant 'distinct' applied to an explicit list as built-in
 
boehmes 
parents: 
40662 
diff
changeset
 | 
165  | 
val distinct = U.mk_const_pat @{theory} @{const_name distinct}
 | 
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
166  | 
(U.destT1 o U.destT1)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
167  | 
fun mk_distinct [] = mk_true  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
168  | 
| mk_distinct (cts as (ct :: _)) =  | 
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
169  | 
Thm.capply (U.instT' ct distinct) (mk_list (Thm.ctyp_of_term ct) cts)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
170  | 
|
| 
41127
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
171  | 
val access = U.mk_const_pat @{theory} @{const_name fun_app} U.destT1
 | 
| 
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
172  | 
fun mk_access array = Thm.capply (U.instT' array access) array  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
173  | 
|
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
174  | 
val update = U.mk_const_pat @{theory} @{const_name fun_upd}
 | 
| 
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
175  | 
(Thm.dest_ctyp o U.destT1)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
176  | 
fun mk_update array index value =  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
177  | 
let val cTs = Thm.dest_ctyp (Thm.ctyp_of_term array)  | 
| 
40662
 
798aad2229c0
added prove reconstruction for injective functions;
 
boehmes 
parents: 
40579 
diff
changeset
 | 
178  | 
in Thm.capply (Thm.mk_binop (U.instTs cTs update) array index) value end  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
179  | 
|
| 
40579
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
180  | 
val mk_uminus = Thm.capply (Thm.cterm_of @{theory} @{const uminus (int)})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
181  | 
val mk_add = Thm.mk_binop (Thm.cterm_of @{theory} @{const plus (int)})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
182  | 
val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (int)})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
183  | 
val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (int)})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
184  | 
val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3div})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
185  | 
val mk_mod = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3mod})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
186  | 
val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (int)})
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
187  | 
val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (int)})
 | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
188  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
189  | 
fun mk_builtin_fun ctxt sym cts =  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
190  | 
(case (sym, cts) of  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
191  | 
    (Sym ("true", _), []) => SOME mk_true
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
192  | 
  | (Sym ("false", _), []) => SOME mk_false
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
193  | 
  | (Sym ("not", _), [ct]) => SOME (mk_not ct)
 | 
| 
40579
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
194  | 
  | (Sym ("and", _), _) => SOME (mk_nary conj mk_true cts)
 | 
| 
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40516 
diff
changeset
 | 
195  | 
  | (Sym ("or", _), _) => SOME (mk_nary disj mk_false cts)
 | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
196  | 
  | (Sym ("implies", _), [ct, cu]) => SOME (mk_implies ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
197  | 
  | (Sym ("iff", _), [ct, cu]) => SOME (mk_iff ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
198  | 
  | (Sym ("~", _), [ct, cu]) => SOME (mk_iff ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
199  | 
  | (Sym ("xor", _), [ct, cu]) => SOME (mk_not (mk_iff ct cu))
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
200  | 
  | (Sym ("ite", _), [ct1, ct2, ct3]) => SOME (mk_if ct1 ct2 ct3)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
201  | 
  | (Sym ("=", _), [ct, cu]) => SOME (mk_eq ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
202  | 
  | (Sym ("distinct", _), _) => SOME (mk_distinct cts)
 | 
| 
41127
 
2ea84c8535c6
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
 
boehmes 
parents: 
41126 
diff
changeset
 | 
203  | 
  | (Sym ("select", _), [ca, ck]) => SOME (Thm.capply (mk_access ca) ck)
 | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
204  | 
  | (Sym ("store", _), [ca, ck, cv]) => SOME (mk_update ca ck cv)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
205  | 
| _ =>  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
206  | 
(case (sym, try (#T o Thm.rep_cterm o hd) cts, cts) of  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
207  | 
      (Sym ("+", _), SOME @{typ int}, [ct, cu]) => SOME (mk_add ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
208  | 
    | (Sym ("-", _), SOME @{typ int}, [ct]) => SOME (mk_uminus ct)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
209  | 
    | (Sym ("-", _), SOME @{typ int}, [ct, cu]) => SOME (mk_sub ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
210  | 
    | (Sym ("*", _), SOME @{typ int}, [ct, cu]) => SOME (mk_mul ct cu)
 | 
| 37151 | 211  | 
    | (Sym ("div", _), SOME @{typ int}, [ct, cu]) => SOME (mk_div ct cu)
 | 
212  | 
    | (Sym ("mod", _), SOME @{typ int}, [ct, cu]) => SOME (mk_mod ct cu)
 | 
|
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
213  | 
    | (Sym ("<", _), SOME @{typ int}, [ct, cu]) => SOME (mk_lt ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
214  | 
    | (Sym ("<=", _), SOME @{typ int}, [ct, cu]) => SOME (mk_le ct cu)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
215  | 
    | (Sym (">", _), SOME @{typ int}, [ct, cu]) => SOME (mk_lt cu ct)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
216  | 
    | (Sym (">=", _), SOME @{typ int}, [ct, cu]) => SOME (mk_le cu ct)
 | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
217  | 
| _ => chained_mk_builtin_fun ctxt (get_mk_builtins ctxt) sym cts))  | 
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
218  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
219  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
220  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
221  | 
(* abstraction *)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
222  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
223  | 
fun is_builtin_theory_term ctxt t =  | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
224  | 
if B.is_builtin_num ctxt t then true  | 
| 
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
225  | 
else  | 
| 
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
226  | 
(case Term.strip_comb t of  | 
| 
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
227  | 
(Const c, ts) => B.is_builtin_fun ctxt c ts  | 
| 
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40681 
diff
changeset
 | 
228  | 
| _ => false)  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
229  | 
|
| 
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
230  | 
end  |