| author | paulson <lp15@cam.ac.uk> | 
| Mon, 27 Feb 2023 17:09:59 +0000 | |
| changeset 77406 | c2013f617a70 | 
| parent 75611 | 66edc020a322 | 
| child 80910 | 406a85a25189 | 
| permissions | -rw-r--r-- | 
| 58061 | 1 | (* Title: HOL/Tools/SMT/smt_translate.ML | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 2 | Author: Sascha Boehme, TU Muenchen | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 3 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 4 | Translate theorems into an SMT intermediate format and serialize them. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 5 | *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 6 | |
| 58061 | 7 | signature SMT_TRANSLATE = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 8 | sig | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 9 | (*intermediate term structure*) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 10 | datatype squant = SForall | SExists | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 11 | datatype 'a spattern = SPat of 'a list | SNoPat of 'a list | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 12 | datatype sterm = | 
| 66551 | 13 | SVar of int * sterm list | | 
| 14 | SConst of string * sterm list | | |
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56811diff
changeset | 15 | SQua of squant * string list * sterm spattern list * sterm | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 16 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 17 | (*translation configuration*) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 18 |   type sign = {
 | 
| 57238 | 19 | logic: string, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 20 | sorts: string list, | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 21 | dtyps: (BNF_Util.fp_kind * (string * (string * (string * string) list) list)) list, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 22 | funcs: (string * (string list * string)) list } | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 23 |   type config = {
 | 
| 66551 | 24 | order: SMT_Util.order, | 
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 25 | logic: string -> term list -> string, | 
| 58360 | 26 | fp_kinds: BNF_Util.fp_kind list, | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 27 | serialize: (string * string) list -> string list -> sign -> (SMT_Util.role * sterm) list -> | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 28 | string } | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 29 |   type replay_data = {
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 30 | context: Proof.context, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 31 | typs: typ Symtab.table, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 32 | terms: term Symtab.table, | 
| 57541 | 33 | ll_defs: term list, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 34 | rewrite_rules: thm list, | 
| 75365 | 35 | assms: ((int * SMT_Util.role) * thm) list } | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 36 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 37 | (*translation*) | 
| 58061 | 38 | val add_config: SMT_Util.class * (Proof.context -> config) -> Context.generic -> Context.generic | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 39 | val translate: Proof.context -> string -> (string * string) list -> string list -> | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 40 | ((int * SMT_Util.role) * thm) list -> | 
| 57239 | 41 | string * replay_data | 
| 57229 | 42 | end; | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 43 | |
| 58061 | 44 | structure SMT_Translate: SMT_TRANSLATE = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 45 | struct | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 46 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 47 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 48 | (* intermediate term structure *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 49 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 50 | datatype squant = SForall | SExists | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 51 | |
| 66134 
a1fb6beb2731
correctly unfold applied 'let's (e.g. '(let x = a in f) b') -- and removed dead code
 blanchet parents: 
61782diff
changeset | 52 | datatype 'a spattern = | 
| 
a1fb6beb2731
correctly unfold applied 'let's (e.g. '(let x = a in f) b') -- and removed dead code
 blanchet parents: 
61782diff
changeset | 53 | SPat of 'a list | SNoPat of 'a list | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 54 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 55 | datatype sterm = | 
| 66551 | 56 | SVar of int * sterm list | | 
| 57 | SConst of string * sterm list | | |
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56811diff
changeset | 58 | SQua of squant * string list * sterm spattern list * sterm | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 59 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 60 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 61 | (* translation configuration *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 63 | type sign = {
 | 
| 57238 | 64 | logic: string, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 65 | sorts: string list, | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 66 | dtyps: (BNF_Util.fp_kind * (string * (string * (string * string) list) list)) list, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 67 | funcs: (string * (string list * string)) list } | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 68 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 69 | type config = {
 | 
| 66551 | 70 | order: SMT_Util.order, | 
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 71 | logic: string -> term list -> string, | 
| 58360 | 72 | fp_kinds: BNF_Util.fp_kind list, | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 73 | serialize: (string * string) list -> string list -> sign -> (SMT_Util.role * sterm) list -> | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 74 | string } | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 75 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 76 | type replay_data = {
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 77 | context: Proof.context, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 78 | typs: typ Symtab.table, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 79 | terms: term Symtab.table, | 
| 57541 | 80 | ll_defs: term list, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 81 | rewrite_rules: thm list, | 
| 75365 | 82 | assms: ((int * SMT_Util.role) * thm) list } | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 83 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 85 | (* translation context *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 86 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 87 | fun add_components_of_typ (Type (s, Ts)) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 88 | cons (Long_Name.base_name s) #> fold_rev add_components_of_typ Ts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 89 | | add_components_of_typ (TFree (s, _)) = cons (perhaps (try (unprefix "'")) s) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 90 | | add_components_of_typ _ = I; | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 91 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 92 | fun suggested_name_of_typ T = space_implode "_" (add_components_of_typ T []); | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 93 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 94 | fun suggested_name_of_term (Const (s, _)) = Long_Name.base_name s | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 95 | | suggested_name_of_term (Free (s, _)) = s | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 96 | | suggested_name_of_term _ = Name.uu | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 97 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 98 | val empty_tr_context = (Name.context, Typtab.empty, Termtab.empty) | 
| 56096 
3e717b56e955
avoid names that may clash with Z3's output (e.g. '')
 blanchet parents: 
56090diff
changeset | 99 | val safe_suffix = "$" | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 100 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 101 | fun add_typ T proper (cx as (names, typs, terms)) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 102 | (case Typtab.lookup typs T of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 103 | SOME (name, _) => (name, cx) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 104 | | NONE => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 105 | let | 
| 56811 | 106 | val sugg = Name.desymbolize (SOME true) (suggested_name_of_typ T) ^ safe_suffix | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 107 | val (name, names') = Name.variant sugg names | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 108 | val typs' = Typtab.update (T, (name, proper)) typs | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 109 | in (name, (names', typs', terms)) end) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 110 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 111 | fun add_fun t sort (cx as (names, typs, terms)) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 112 | (case Termtab.lookup terms t of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 113 | SOME (name, _) => (name, cx) | 
| 57230 | 114 | | NONE => | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 115 | let | 
| 56811 | 116 | val sugg = Name.desymbolize (SOME false) (suggested_name_of_term t) ^ safe_suffix | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 117 | val (name, names') = Name.variant sugg names | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 118 | val terms' = Termtab.update (t, (name, sort)) terms | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 119 | in (name, (names', typs, terms')) end) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 120 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 121 | fun sign_of logic dtyps (_, typs, terms) = {
 | 
| 57238 | 122 | logic = logic, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 123 | sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [], | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 124 | dtyps = dtyps, | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 125 | funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms []} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 126 | |
| 57541 | 127 | fun replay_data_of ctxt ll_defs rules assms (_, typs, terms) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 128 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 129 | fun add_typ (T, (n, _)) = Symtab.update (n, T) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 130 | val typs' = Typtab.fold add_typ typs Symtab.empty | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 131 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 132 | fun add_fun (t, (n, _)) = Symtab.update (n, t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 133 | val terms' = Termtab.fold add_fun terms Symtab.empty | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 134 | in | 
| 57541 | 135 |     {context = ctxt, typs = typs', terms = terms', ll_defs = ll_defs, rewrite_rules = rules,
 | 
| 75365 | 136 | assms = assms} | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 137 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 138 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 139 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 140 | (* preprocessing *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 141 | |
| 58361 | 142 | (** (co)datatype declarations **) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 143 | |
| 58361 | 144 | fun collect_co_datatypes fp_kinds (tr_context, ctxt) ts = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 145 | let | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 146 | val (fp_decls, ctxt') = | 
| 58361 | 147 | ([], ctxt) | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 148 | |> fold (Term.fold_types (SMT_Datatypes.add_decls fp_kinds) o snd) ts | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 149 | |>> flat | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 150 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 151 | fun is_decl_typ T = exists (equal T o fst o snd) fp_decls | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 152 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 153 | fun add_typ' T proper = | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 154 | (case SMT_Builtin.dest_builtin_typ ctxt' T of | 
| 66551 | 155 | SOME (n, Ts) => pair n (* FIXME HO: Consider Ts *) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 156 | | NONE => add_typ T proper) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 157 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 158 | fun tr_select sel = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 159 | let val T = Term.range_type (Term.fastype_of sel) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 160 | in add_fun sel NONE ##>> add_typ' T (not (is_decl_typ T)) end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 161 | fun tr_constr (constr, selects) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 162 | add_fun constr NONE ##>> fold_map tr_select selects | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 163 | fun tr_typ (fp, (T, cases)) = | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 164 | add_typ' T false ##>> fold_map tr_constr cases #>> pair fp | 
| 58361 | 165 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 166 | val (fp_decls', tr_context') = fold_map tr_typ fp_decls tr_context | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 167 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 168 | fun add (constr, selects) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 169 | Termtab.update (constr, length selects) #> | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 170 | fold (Termtab.update o rpair 1) selects | 
| 58361 | 171 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 172 | val funcs = fold (fold add o snd o snd) fp_decls Termtab.empty | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 173 | |
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 174 | in ((funcs, fp_decls', tr_context', ctxt'), ts) end | 
| 58361 | 175 | (* FIXME: also return necessary (co)datatype theorems *) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 176 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 177 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 178 | (** eta-expand quantifiers, let expressions and built-ins *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 179 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 180 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 181 | fun eta f T t = Abs (Name.uu, T, f (Term.incr_boundvars 1 t $ Bound 0)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 182 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 183 | fun exp f T = eta f (Term.domain_type (Term.domain_type T)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 184 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 185 | fun exp2 T q = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 186 | let val U = Term.domain_type T | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 187 | in Abs (Name.uu, U, q $ eta I (Term.domain_type U) (Bound 0)) end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 188 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 189 | fun expf k i T t = | 
| 58061 | 190 | let val Ts = drop i (fst (SMT_Util.dest_funT k T)) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 191 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 192 | Term.incr_boundvars (length Ts) t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 193 | |> fold_rev (fn i => fn u => u $ Bound i) (0 upto length Ts - 1) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 194 | |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 195 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 196 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 197 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 198 | fun eta_expand ctxt funcs = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 199 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 200 | fun exp_func t T ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 201 | (case Termtab.lookup funcs t of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 202 | SOME k => Term.list_comb (t, ts) |> k <> length ts ? expf k (length ts) T | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 203 | | NONE => Term.list_comb (t, ts)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 204 | |
| 67149 | 205 | fun expand ((q as Const (\<^const_name>\<open>All\<close>, _)) $ Abs a) = q $ abs_expand a | 
| 206 | | expand ((q as Const (\<^const_name>\<open>All\<close>, T)) $ t) = q $ exp expand T t | |
| 207 | | expand (q as Const (\<^const_name>\<open>All\<close>, T)) = exp2 T q | |
| 208 | | expand ((q as Const (\<^const_name>\<open>Ex\<close>, _)) $ Abs a) = q $ abs_expand a | |
| 209 | | expand ((q as Const (\<^const_name>\<open>Ex\<close>, T)) $ t) = q $ exp expand T t | |
| 210 | | expand (q as Const (\<^const_name>\<open>Ex\<close>, T)) = exp2 T q | |
| 211 | | expand (Const (\<^const_name>\<open>Let\<close>, T) $ t) = | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 212 | let val U = Term.domain_type (Term.range_type T) | 
| 75611 
66edc020a322
missing recursive let-expansion in SMT translation
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75365diff
changeset | 213 | in Abs (Name.uu, U, expand (Bound 0 $ Term.incr_boundvars 1 t)) end | 
| 67149 | 214 | | expand (Const (\<^const_name>\<open>Let\<close>, T)) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 215 | let val U = Term.domain_type (Term.range_type T) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 216 | in Abs (Name.uu, Term.domain_type T, Abs (Name.uu, U, Bound 0 $ Bound 1)) end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 217 | | expand t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 218 | (case Term.strip_comb t of | 
| 67149 | 219 | (Const (\<^const_name>\<open>Let\<close>, _), t1 :: t2 :: ts) => | 
| 66136 | 220 | Term.betapplys (Term.betapply (expand t2, expand t1), map expand ts) | 
| 66134 
a1fb6beb2731
correctly unfold applied 'let's (e.g. '(let x = a in f) b') -- and removed dead code
 blanchet parents: 
61782diff
changeset | 221 | | (u as Const (c as (_, T)), ts) => | 
| 58061 | 222 | (case SMT_Builtin.dest_builtin ctxt c ts of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 223 | SOME (_, k, us, mk) => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 224 | if k = length us then mk (map expand us) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 225 | else if k < length us then chop k (map expand us) |>> mk |> Term.list_comb | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 226 | else expf k (length ts) T (mk (map expand us)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 227 | | NONE => exp_func u T (map expand ts)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 228 | | (u as Free (_, T), ts) => exp_func u T (map expand ts) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 229 | | (Abs a, ts) => Term.list_comb (abs_expand a, map expand ts) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 230 | | (u, ts) => Term.list_comb (u, map expand ts)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 231 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 232 | and abs_expand (n, T, t) = Abs (n, T, expand t) | 
| 57230 | 233 | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 234 | in map (apsnd expand) end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 235 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 236 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 237 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 238 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 239 | (** introduce explicit applications **) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 240 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 241 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 242 | (* | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 243 | Make application explicit for functions with varying number of arguments. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 244 | *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 245 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 246 | fun add t i = apfst (Termtab.map_default (t, i) (Integer.min i)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 247 | fun add_type T = apsnd (Typtab.update (T, ())) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 248 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 249 | fun min_arities t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 250 | (case Term.strip_comb t of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 251 | (u as Const _, ts) => add u (length ts) #> fold min_arities ts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 252 | | (u as Free _, ts) => add u (length ts) #> fold min_arities ts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 253 | | (Abs (_, T, u), ts) => (can dest_funT T ? add_type T) #> min_arities u #> fold min_arities ts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 254 | | (_, ts) => fold min_arities ts) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 255 | |
| 66738 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 256 | fun take_vars_into_account types t i = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 257 | let | 
| 67149 | 258 | fun find_min j (T as Type (\<^type_name>\<open>fun\<close>, [_, T'])) = | 
| 66738 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 259 | if j = i orelse Typtab.defined types T then j else find_min (j + 1) T' | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 260 | | find_min j _ = j | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 261 | in find_min 0 (Term.type_of t) end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 262 | |
| 67149 | 263 | fun app u (t, T) = (Const (\<^const_name>\<open>fun_app\<close>, T --> T) $ t $ u, Term.range_type T) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 264 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 265 | fun apply i t T ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 266 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 267 | val (ts1, ts2) = chop i ts | 
| 58061 | 268 | val (_, U) = SMT_Util.dest_funT i T | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 269 | in fst (fold app ts2 (Term.list_comb (t, ts1), U)) end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 270 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 271 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 272 | fun intro_explicit_application ctxt funcs ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 273 | let | 
| 66738 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 274 | val explicit_application = Config.get ctxt SMT_Config.explicit_application | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 275 | val get_arities = | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 276 | (case explicit_application of | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 277 | 0 => min_arities | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 278 | | 1 => min_arities | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 279 | | 2 => K I | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 280 |       | n => error ("Illegal value for " ^ quote (Config.name_of SMT_Config.explicit_application) ^
 | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 281 | ": " ^ string_of_int n)) | 
| 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 282 | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 283 | val (arities, types) = fold (get_arities o snd) ts (Termtab.empty, Typtab.empty) | 
| 66738 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 284 | val arities' = arities |> explicit_application = 1 ? Termtab.map (take_vars_into_account types) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 285 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 286 | fun app_func t T ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 287 | if is_some (Termtab.lookup funcs t) then Term.list_comb (t, ts) | 
| 66738 
793e7a9c30c5
properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
 blanchet parents: 
66551diff
changeset | 288 | else apply (the_default 0 (Termtab.lookup arities' t)) t T ts | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 289 | |
| 58061 | 290 | fun in_list T f t = SMT_Util.mk_symb_list T (map f (SMT_Util.dest_symb_list t)) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 291 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 292 | fun traverse Ts t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 293 | (case Term.strip_comb t of | 
| 67149 | 294 | (q as Const (\<^const_name>\<open>All\<close>, _), [Abs (x, T, u)]) => | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 295 | q $ Abs (x, T, in_trigger (T :: Ts) u) | 
| 67149 | 296 | | (q as Const (\<^const_name>\<open>Ex\<close>, _), [Abs (x, T, u)]) => | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 297 | q $ Abs (x, T, in_trigger (T :: Ts) u) | 
| 67149 | 298 | | (q as Const (\<^const_name>\<open>Let\<close>, _), [u1, u2 as Abs _]) => | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 299 | q $ traverse Ts u1 $ traverse Ts u2 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 300 | | (u as Const (c as (_, T)), ts) => | 
| 58061 | 301 | (case SMT_Builtin.dest_builtin ctxt c ts of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 302 | SOME (_, k, us, mk) => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 303 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 304 | val (ts1, ts2) = chop k (map (traverse Ts) us) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 305 | val U = Term.strip_type T |>> snd o chop k |> (op --->) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 306 | in apply 0 (mk ts1) U ts2 end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 307 | | NONE => app_func u T (map (traverse Ts) ts)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 308 | | (u as Free (_, T), ts) => app_func u T (map (traverse Ts) ts) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 309 | | (u as Bound i, ts) => apply 0 u (nth Ts i) (map (traverse Ts) ts) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 310 | | (Abs (n, T, u), ts) => traverses Ts (Abs (n, T, traverse (T::Ts) u)) ts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 311 | | (u, ts) => traverses Ts u ts) | 
| 74382 | 312 | and in_trigger Ts ((c as \<^Const_>\<open>trigger\<close>) $ p $ t) = c $ in_pats Ts p $ traverse Ts t | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56811diff
changeset | 313 | | in_trigger Ts t = traverse Ts t | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 314 | and in_pats Ts ps = | 
| 67149 | 315 | in_list \<^typ>\<open>pattern symb_list\<close> (in_list \<^typ>\<open>pattern\<close> (in_pat Ts)) ps | 
| 74382 | 316 | and in_pat Ts ((p as \<^Const_>\<open>pat _\<close>) $ t) = p $ traverse Ts t | 
| 317 | | in_pat Ts ((p as \<^Const_>\<open>nopat _\<close>) $ t) = p $ traverse Ts t | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 318 |       | in_pat _ t = raise TERM ("bad pattern", [t])
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 319 | and traverses Ts t ts = Term.list_comb (t, map (traverse Ts) ts) | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 320 | in map (apsnd (traverse [])) ts end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 321 | |
| 57230 | 322 | val fun_app_eq = mk_meta_eq @{thm fun_app_def}
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 323 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 324 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 325 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 326 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 327 | (** map HOL formulas to FOL formulas (i.e., separate formulas froms terms) **) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 328 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 329 | local | 
| 67149 | 330 | val is_quant = member (op =) [\<^const_name>\<open>All\<close>, \<^const_name>\<open>Ex\<close>] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 331 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 332 | val fol_rules = [ | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 333 | Let_def, | 
| 61782 
7d754aca6a75
removed needless complication for modern SMT solvers
 blanchet parents: 
58429diff
changeset | 334 |     @{lemma "P = True == P" by (rule eq_reflection) simp}]
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 335 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 336 | exception BAD_PATTERN of unit | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 337 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 338 | fun is_builtin_conn_or_pred ctxt c ts = | 
| 58061 | 339 | is_some (SMT_Builtin.dest_builtin_conn ctxt c ts) orelse | 
| 340 | is_some (SMT_Builtin.dest_builtin_pred ctxt c ts) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 341 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 342 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 343 | fun folify ctxt = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 344 | let | 
| 58061 | 345 | fun in_list T f t = SMT_Util.mk_symb_list T (map_filter f (SMT_Util.dest_symb_list t)) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 346 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 347 | fun in_term pat t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 348 | (case Term.strip_comb t of | 
| 74382 | 349 | (\<^Const_>\<open>True\<close>, []) => t | 
| 350 | | (\<^Const_>\<open>False\<close>, []) => t | |
| 351 | | (u as \<^Const_>\<open>If _\<close>, [t1, t2, t3]) => | |
| 56104 
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
 blanchet parents: 
56096diff
changeset | 352 | if pat then raise BAD_PATTERN () else u $ in_form t1 $ in_term pat t2 $ in_term pat t3 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 353 | | (Const (c as (n, _)), ts) => | 
| 61782 
7d754aca6a75
removed needless complication for modern SMT solvers
 blanchet parents: 
58429diff
changeset | 354 | if is_builtin_conn_or_pred ctxt c ts orelse is_quant n then | 
| 
7d754aca6a75
removed needless complication for modern SMT solvers
 blanchet parents: 
58429diff
changeset | 355 | if pat then raise BAD_PATTERN () else in_form t | 
| 
7d754aca6a75
removed needless complication for modern SMT solvers
 blanchet parents: 
58429diff
changeset | 356 | else | 
| 
7d754aca6a75
removed needless complication for modern SMT solvers
 blanchet parents: 
58429diff
changeset | 357 | Term.list_comb (Const c, map (in_term pat) ts) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 358 | | (Free c, ts) => Term.list_comb (Free c, map (in_term pat) ts) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 359 | | _ => t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 360 | |
| 67149 | 361 | and in_pat ((p as Const (\<^const_name>\<open>pat\<close>, _)) $ t) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 362 | p $ in_term true t | 
| 67149 | 363 | | in_pat ((p as Const (\<^const_name>\<open>nopat\<close>, _)) $ t) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 364 | p $ in_term true t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 365 |       | in_pat t = raise TERM ("bad pattern", [t])
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 366 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 367 | and in_pats ps = | 
| 67149 | 368 | in_list \<^typ>\<open>pattern symb_list\<close> (SOME o in_list \<^typ>\<open>pattern\<close> (try in_pat)) ps | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 369 | |
| 74382 | 370 | and in_trigger ((c as \<^Const_>\<open>trigger\<close>) $ p $ t) = c $ in_pats p $ in_form t | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56811diff
changeset | 371 | | in_trigger t = in_form t | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 372 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 373 | and in_form t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 374 | (case Term.strip_comb t of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 375 | (q as Const (qn, _), [Abs (n, T, u)]) => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 376 | if is_quant qn then q $ Abs (n, T, in_trigger u) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 377 | else in_term false t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 378 | | (Const c, ts) => | 
| 58061 | 379 | (case SMT_Builtin.dest_builtin_conn ctxt c ts of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 380 | SOME (_, _, us, mk) => mk (map in_form us) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 381 | | NONE => | 
| 58061 | 382 | (case SMT_Builtin.dest_builtin_pred ctxt c ts of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 383 | SOME (_, _, us, mk) => mk (map (in_term false) us) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 384 | | NONE => in_term false t)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 385 | | _ => in_term false t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 386 | in | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 387 | map (apsnd in_form) #> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 388 | pair (fol_rules, I) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 389 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 390 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 391 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 392 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 393 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 394 | (* translation into intermediate format *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 395 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 396 | (** utility functions **) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 397 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 398 | val quantifier = (fn | 
| 67149 | 399 | \<^const_name>\<open>All\<close> => SOME SForall | 
| 400 | | \<^const_name>\<open>Ex\<close> => SOME SExists | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 401 | | _ => NONE) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 402 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 403 | fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 404 | if q = qname then group_quant qname (T :: Ts) u else (Ts, t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 405 | | group_quant _ Ts t = (Ts, t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 406 | |
| 67149 | 407 | fun dest_pat (Const (\<^const_name>\<open>pat\<close>, _) $ t) = (t, true) | 
| 408 | | dest_pat (Const (\<^const_name>\<open>nopat\<close>, _) $ t) = (t, false) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 409 |   | dest_pat t = raise TERM ("bad pattern", [t])
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 410 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 411 | fun dest_pats [] = I | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 412 | | dest_pats ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 413 | (case map dest_pat ts |> split_list ||> distinct (op =) of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 414 | (ps, [true]) => cons (SPat ps) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 415 | | (ps, [false]) => cons (SNoPat ps) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 416 |       | _ => raise TERM ("bad multi-pattern", ts))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 417 | |
| 74382 | 418 | fun dest_trigger \<^Const_>\<open>trigger for tl t\<close> = | 
| 58061 | 419 | (rev (fold (dest_pats o SMT_Util.dest_symb_list) (SMT_Util.dest_symb_list tl) []), t) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 420 | | dest_trigger t = ([], t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 421 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 422 | fun dest_quant qn T t = quantifier qn |> Option.map (fn q => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 423 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 424 | val (Ts, u) = group_quant qn [T] t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 425 | val (ps, p) = dest_trigger u | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56811diff
changeset | 426 | in (q, rev Ts, ps, p) end) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 427 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 428 | fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 429 | | fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 430 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 431 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 432 | (** translation from Isabelle terms into SMT intermediate terms **) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 433 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 434 | fun intermediate logic dtyps builtin ctxt ts trx = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 435 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 436 | fun transT (T as TFree _) = add_typ T true | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 437 |       | transT (T as TVar _) = (fn _ => raise TYPE ("bad SMT type", [T], []))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 438 | | transT (T as Type _) = | 
| 58061 | 439 | (case SMT_Builtin.dest_builtin_typ ctxt T of | 
| 66551 | 440 | SOME (n, []) => pair n | 
| 441 | | SOME (n, Ts) => | |
| 442 | fold_map transT Ts | |
| 443 |             #>> (fn ns => enclose "(" ")" (space_implode " " (n :: ns)))
 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 444 | | NONE => add_typ T true) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 445 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 446 | fun trans t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 447 | (case Term.strip_comb t of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 448 | (Const (qn, _), [Abs (_, T, t1)]) => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 449 | (case dest_quant qn T t1 of | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56811diff
changeset | 450 | SOME (q, Ts, ps, b) => | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 451 | fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>> | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56811diff
changeset | 452 | trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', b')) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 453 |           | NONE => raise TERM ("unsupported quantifier", [t]))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 454 | | (u as Const (c as (_, T)), ts) => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 455 | (case builtin ctxt c ts of | 
| 66551 | 456 | SOME (n, _, us, _) => fold_map trans us #>> curry SConst n | 
| 457 | | NONE => trans_applied_fun u T ts) | |
| 458 | | (u as Free (_, T), ts) => trans_applied_fun u T ts | |
| 459 | | (Bound i, ts) => pair i ##>> fold_map trans ts #>> SVar | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 460 |       | _ => raise TERM ("bad SMT term", [t]))
 | 
| 57230 | 461 | |
| 66551 | 462 | and trans_applied_fun t T ts = | 
| 58061 | 463 | let val (Us, U) = SMT_Util.dest_funT (length ts) T | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 464 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 465 | fold_map transT Us ##>> transT U #-> (fn Up => | 
| 66551 | 466 | add_fun t (SOME Up) ##>> fold_map trans ts #>> SConst) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 467 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 468 | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 469 | val (us, trx') = fold_map (fn (role, t) => apfst (pair role) o trans t) ts trx | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 470 | in ((sign_of (logic ts) dtyps trx', us), trx') end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 471 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 472 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 473 | (* translation *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 474 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 475 | structure Configs = Generic_Data | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 476 | ( | 
| 58061 | 477 | type T = (Proof.context -> config) SMT_Util.dict | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 478 | val empty = [] | 
| 58061 | 479 | fun merge data = SMT_Util.dict_merge fst data | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 480 | ) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 481 | |
| 58061 | 482 | fun add_config (cs, cfg) = Configs.map (SMT_Util.dict_update (cs, cfg)) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 483 | |
| 57230 | 484 | fun get_config ctxt = | 
| 58061 | 485 | let val cs = SMT_Config.solver_class_of ctxt | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 486 | in | 
| 58061 | 487 | (case SMT_Util.dict_get (Configs.get (Context.Proof ctxt)) cs of | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 488 | SOME cfg => cfg ctxt | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 489 |     | NONE => error ("SMT: no translation configuration found " ^
 | 
| 58061 | 490 | "for solver class " ^ quote (SMT_Util.string_of_class cs))) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 491 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 492 | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 493 | fun translate ctxt prover smt_options comments (ithms : ((int * SMT_Util.role) * thm) list) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 494 | let | 
| 66551 | 495 |     val {order, logic, fp_kinds, serialize} = get_config ctxt
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 496 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 497 | fun no_dtyps (tr_context, ctxt) ts = | 
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 498 | ((Termtab.empty, [], tr_context, ctxt), ts) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 499 | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 500 | val ts1 = map (SMT_Util.map_prod snd (Envir.beta_eta_contract o SMT_Util.prop_of)) ithms | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 501 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 502 | val ((funcs, dtyps, tr_context, ctxt1), ts2) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 503 | ((empty_tr_context, ctxt), ts1) | 
| 58361 | 504 | |-> (if null fp_kinds then no_dtyps else collect_co_datatypes fp_kinds) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 505 | |
| 67149 | 506 | fun is_binder (Const (\<^const_name>\<open>Let\<close>, _) $ _) = true | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 507 | | is_binder t = Lambda_Lifting.is_quantifier t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 508 | |
| 67149 | 509 | fun mk_trigger ((q as Const (\<^const_name>\<open>All\<close>, _)) $ Abs (n, T, t)) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 510 | q $ Abs (n, T, mk_trigger t) | 
| 67149 | 511 | | mk_trigger (eq as (Const (\<^const_name>\<open>HOL.eq\<close>, T) $ lhs $ _)) = | 
| 512 | Term.domain_type T --> \<^typ>\<open>pattern\<close> | |
| 513 | |> (fn T => Const (\<^const_name>\<open>pat\<close>, T) $ lhs) | |
| 514 | |> SMT_Util.mk_symb_list \<^typ>\<open>pattern\<close> o single | |
| 515 | |> SMT_Util.mk_symb_list \<^typ>\<open>pattern symb_list\<close> o single | |
| 74382 | 516 | |> (fn t => \<^Const>\<open>trigger for t eq\<close>) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 517 | | mk_trigger t = t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 518 | |
| 57541 | 519 | val (ctxt2, (ts3, ll_defs)) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 520 | ts2 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 521 | |> eta_expand ctxt1 funcs | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 522 | |> rpair ctxt1 | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 523 | |-> Lambda_Lifting.lift_lambdas' NONE is_binder | 
| 57541 | 524 | |-> (fn (ts', ll_defs) => fn ctxt' => | 
| 66551 | 525 | let | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 526 | val ts'' = map (pair SMT_Util.Axiom o mk_trigger) ll_defs @ ts' | 
| 66551 | 527 | |> order = SMT_Util.First_Order ? intro_explicit_application ctxt' funcs | 
| 528 | in | |
| 529 | (ctxt', (ts'', ll_defs)) | |
| 530 | end) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 531 | val ((rewrite_rules, builtin), ts4) = folify ctxt2 ts3 | 
| 66551 | 532 | |>> order = SMT_Util.First_Order ? apfst (cons fun_app_eq) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 533 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 534 | (ts4, tr_context) | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 535 | |-> intermediate (logic prover o map snd) dtyps (builtin SMT_Builtin.dest_builtin) ctxt2 | 
| 57239 | 536 | |>> uncurry (serialize smt_options comments) | 
| 57541 | 537 | ||> replay_data_of ctxt2 ll_defs rewrite_rules ithms | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 538 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 539 | |
| 57229 | 540 | end; |