src/HOL/Tools/SMT/smt_translate.ML
author blanchet
Fri, 04 Dec 2015 14:15:16 +0100
changeset 61782 7d754aca6a75
parent 58429 0b94858325a5
child 66134 a1fb6beb2731
permissions -rw-r--r--
removed needless complication for modern SMT solvers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
     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
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
     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 =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    13
    SVar of int |
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    14
    SApp of string * sterm list |
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
    SLet of string * sterm * sterm |
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
    16
    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
    17
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    18
  (*translation configuration*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
  type sign = {
57238
blanchet
parents: 57230
diff changeset
    20
    logic: string,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    21
    sorts: string list,
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
    22
    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
    23
    funcs: (string * (string list * string)) list }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    24
  type config = {
57238
blanchet
parents: 57230
diff changeset
    25
    logic: term list -> string,
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
    26
    fp_kinds: BNF_Util.fp_kind list,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57238
diff changeset
    27
    serialize: (string * string) list -> string list -> sign -> sterm list -> string }
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
  type replay_data = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    29
    context: Proof.context,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    30
    typs: typ Symtab.table,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
    terms: term Symtab.table,
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
    32
    ll_defs: term list,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
    rewrite_rules: thm list,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    34
    assms: (int * thm) list }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    36
  (*translation*)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
    37
  val add_config: SMT_Util.class * (Proof.context -> config) -> Context.generic -> Context.generic
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57238
diff changeset
    38
  val translate: Proof.context -> (string * string) list -> string list -> (int * thm) list ->
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57238
diff changeset
    39
    string * replay_data
57229
blanchet
parents: 57165
diff changeset
    40
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    41
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
    42
structure SMT_Translate: SMT_TRANSLATE =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    44
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    45
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    46
(* intermediate term structure *)
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
datatype squant = SForall | SExists
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 '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
    51
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
datatype sterm =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
  SVar of int |
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    54
  SApp of string * sterm list |
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    55
  SLet of string * sterm * sterm |
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
    56
  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
    57
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
(* translation configuration *)
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
type sign = {
57238
blanchet
parents: 57230
diff changeset
    62
  logic: string,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    63
  sorts: string list,
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
    64
  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
    65
  funcs: (string * (string list * string)) list }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    67
type config = {
57238
blanchet
parents: 57230
diff changeset
    68
  logic: term list -> string,
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
    69
  fp_kinds: BNF_Util.fp_kind list,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57238
diff changeset
    70
  serialize: (string * string) list -> string list -> sign -> sterm list -> string }
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    71
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    72
type replay_data = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
  context: Proof.context,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    74
  typs: typ Symtab.table,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    75
  terms: term Symtab.table,
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
    76
  ll_defs: term list,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    77
  rewrite_rules: thm list,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    78
  assms: (int * thm) list }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    79
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    80
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    81
(* translation context *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    82
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    83
fun add_components_of_typ (Type (s, Ts)) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    84
    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
    85
  | 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
    86
  | add_components_of_typ _ = I;
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    87
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
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
    89
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    90
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
    91
  | suggested_name_of_term (Free (s, _)) = s
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    92
  | suggested_name_of_term _ = Name.uu
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
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: 56090
diff changeset
    95
val safe_suffix = "$"
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    96
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    97
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
    98
  (case Typtab.lookup typs T of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    99
    SOME (name, _) => (name, cx)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   100
  | NONE =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   101
      let
56811
b66639331db5 optional case enforcement
haftmann
parents: 56125
diff changeset
   102
        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
   103
        val (name, names') = Name.variant sugg names
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   104
        val typs' = Typtab.update (T, (name, proper)) typs
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
      in (name, (names', typs', terms)) end)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
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
   108
  (case Termtab.lookup terms t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   109
    SOME (name, _) => (name, cx)
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   110
  | NONE =>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   111
      let
56811
b66639331db5 optional case enforcement
haftmann
parents: 56125
diff changeset
   112
        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
   113
        val (name, names') = Name.variant sugg names
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   114
        val terms' = Termtab.update (t, (name, sort)) terms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   115
      in (name, (names', typs, terms')) end)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   116
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   117
fun sign_of logic dtyps (_, typs, terms) = {
57238
blanchet
parents: 57230
diff changeset
   118
  logic = logic,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   119
  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: 58361
diff changeset
   120
  dtyps = dtyps,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   121
  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
   122
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
   123
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
   124
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   125
    fun add_typ (T, (n, _)) = Symtab.update (n, T)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   126
    val typs' = Typtab.fold add_typ typs Symtab.empty
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   127
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   128
    fun add_fun (t, (n, _)) = Symtab.update (n, t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   129
    val terms' = Termtab.fold add_fun terms Symtab.empty
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   130
  in
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
   131
    {context = ctxt, typs = typs', terms = terms', ll_defs = ll_defs, rewrite_rules = rules,
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
   132
     assms = assms}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   133
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   134
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   135
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   136
(* preprocessing *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   137
58361
7f2b3b6f6ad1 added codatatype support for CVC4
blanchet
parents: 58360
diff changeset
   138
(** (co)datatype declarations **)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   139
58361
7f2b3b6f6ad1 added codatatype support for CVC4
blanchet
parents: 58360
diff changeset
   140
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
   141
  let
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   142
    val (fp_decls, ctxt') =
58361
7f2b3b6f6ad1 added codatatype support for CVC4
blanchet
parents: 58360
diff changeset
   143
      ([], ctxt)
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   144
      |> fold (Term.fold_types (SMT_Datatypes.add_decls fp_kinds)) ts
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   145
      |>> flat
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   146
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   147
    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
   148
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   149
    fun add_typ' T proper =
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   150
      (case SMT_Builtin.dest_builtin_typ ctxt' T of
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   151
        SOME n => pair n
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   152
      | NONE => add_typ T proper)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   153
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   154
    fun tr_select sel =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   155
      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
   156
      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
   157
    fun tr_constr (constr, selects) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   158
      add_fun constr NONE ##>> fold_map tr_select selects
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   159
    fun tr_typ (fp, (T, cases)) =
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   160
      add_typ' T false ##>> fold_map tr_constr cases #>> pair fp
58361
7f2b3b6f6ad1 added codatatype support for CVC4
blanchet
parents: 58360
diff changeset
   161
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   162
    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
   163
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   164
    fun add (constr, selects) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
      Termtab.update (constr, length selects) #>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
      fold (Termtab.update o rpair 1) selects
58361
7f2b3b6f6ad1 added codatatype support for CVC4
blanchet
parents: 58360
diff changeset
   167
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   168
    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: 58361
diff changeset
   169
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   170
  in ((funcs, fp_decls', tr_context', ctxt'), ts) end
58361
7f2b3b6f6ad1 added codatatype support for CVC4
blanchet
parents: 58360
diff changeset
   171
    (* FIXME: also return necessary (co)datatype theorems *)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   172
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   173
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   174
(** eta-expand quantifiers, let expressions and built-ins *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   175
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   176
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   177
  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
   178
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   179
  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
   180
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   181
  fun exp2 T q =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   182
    let val U = Term.domain_type T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   183
    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
   184
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   185
  fun expf k i T t =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   186
    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
   187
    in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   188
      Term.incr_boundvars (length Ts) t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   189
      |> 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
   190
      |> 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
   191
    end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   192
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   193
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   194
fun eta_expand ctxt funcs =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   195
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   196
    fun exp_func t T ts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   197
      (case Termtab.lookup funcs t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   198
        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
   199
      | NONE => Term.list_comb (t, ts))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   200
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   201
    fun expand ((q as Const (@{const_name All}, _)) $ Abs a) = q $ abs_expand a
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   202
      | expand ((q as Const (@{const_name All}, T)) $ t) = q $ exp expand T t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   203
      | expand (q as Const (@{const_name All}, T)) = exp2 T q
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   204
      | expand ((q as Const (@{const_name Ex}, _)) $ Abs a) = q $ abs_expand a
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   205
      | expand ((q as Const (@{const_name Ex}, T)) $ t) = q $ exp expand T t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   206
      | expand (q as Const (@{const_name Ex}, T)) = exp2 T q
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
   207
      | expand (Const (@{const_name Let}, _) $ t $ u) = expand (Term.betapply (u, t))
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
   208
      | expand (Const (@{const_name Let}, T) $ t) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   209
          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
   210
          in Abs (Name.uu, U, Bound 0 $ Term.incr_boundvars 1 t) end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   211
      | expand (Const (@{const_name Let}, T)) =
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)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   213
          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
   214
      | expand t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   215
          (case Term.strip_comb t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   216
            (u as Const (c as (_, T)), ts) =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   217
              (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
   218
                SOME (_, k, us, mk) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   219
                  if k = length us then mk (map expand us)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   220
                  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
   221
                  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
   222
              | NONE => exp_func u T (map expand ts))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   223
          | (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
   224
          | (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
   225
          | (u, ts) => Term.list_comb (u, map expand ts))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   226
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   227
    and abs_expand (n, T, t) = Abs (n, T, expand t)
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   228
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   229
  in map expand end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   230
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   231
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   232
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   233
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   234
(** introduce explicit applications **)
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
local
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
    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
   239
  *)
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
  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
   242
  fun add_type T = apsnd (Typtab.update (T, ()))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   243
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   244
  fun min_arities t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   245
    (case Term.strip_comb t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   246
      (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
   247
    | (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
   248
    | (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
   249
    | (_, ts) => fold min_arities ts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   250
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   251
  fun minimize types t i =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   252
    let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   253
      fun find_min j [] _ = j
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   254
        | find_min j (U :: Us) T =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   255
            if Typtab.defined types T then j else find_min (j + 1) Us (U --> T)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   256
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   257
      val (Ts, T) = Term.strip_type (Term.type_of t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   258
    in find_min 0 (take i (rev Ts)) T end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   259
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   260
  fun app u (t, T) = (Const (@{const_name fun_app}, T --> T) $ t $ u, Term.range_type T)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   261
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   262
  fun apply i t T ts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   263
    let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   264
      val (ts1, ts2) = chop i ts
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   265
      val (_, U) = SMT_Util.dest_funT i T
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   266
    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
   267
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   268
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   269
fun intro_explicit_application ctxt funcs ts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   270
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   271
    val (arities, types) = fold min_arities ts (Termtab.empty, Typtab.empty)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   272
    val arities' = Termtab.map (minimize types) arities (* FIXME: highly suspicious *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   273
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   274
    fun app_func t T ts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   275
      if is_some (Termtab.lookup funcs t) then Term.list_comb (t, ts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   276
      else apply (the (Termtab.lookup arities' t)) t T ts
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   277
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   278
    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
   279
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   280
    fun traverse Ts t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   281
      (case Term.strip_comb t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   282
        (q as Const (@{const_name All}, _), [Abs (x, T, u)]) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   283
          q $ Abs (x, T, in_trigger (T :: Ts) u)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   284
      | (q as Const (@{const_name Ex}, _), [Abs (x, T, u)]) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   285
          q $ Abs (x, T, in_trigger (T :: Ts) u)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   286
      | (q as Const (@{const_name Let}, _), [u1, u2 as Abs _]) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   287
          q $ traverse Ts u1 $ traverse Ts u2
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   288
      | (u as Const (c as (_, T)), ts) =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   289
          (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
   290
            SOME (_, k, us, mk) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   291
              let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   292
                val (ts1, ts2) = chop k (map (traverse Ts) us)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   293
                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
   294
              in apply 0 (mk ts1) U ts2 end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   295
          | NONE => app_func u T (map (traverse Ts) ts))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   296
      | (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
   297
      | (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
   298
      | (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
   299
      | (u, ts) => traverses Ts u ts)
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   300
    and in_trigger Ts ((c as @{const trigger}) $ p $ t) = c $ in_pats Ts p $ traverse Ts t
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
   301
      | in_trigger Ts t = traverse Ts t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   302
    and in_pats Ts ps =
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   303
      in_list @{typ "pattern symb_list"} (in_list @{typ pattern} (in_pat Ts)) ps
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   304
    and in_pat Ts ((p as Const (@{const_name pat}, _)) $ t) = p $ traverse Ts t
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   305
      | in_pat Ts ((p as Const (@{const_name nopat}, _)) $ t) = p $ traverse Ts t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   306
      | in_pat _ t = raise TERM ("bad pattern", [t])
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   307
    and traverses Ts t ts = Term.list_comb (t, map (traverse Ts) ts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   308
  in map (traverse []) ts end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   309
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   310
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
   311
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   312
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   313
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   314
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   315
(** 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
   316
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   317
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   318
  val is_quant = member (op =) [@{const_name All}, @{const_name Ex}]
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   319
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   320
  val fol_rules = [
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   321
    Let_def,
61782
7d754aca6a75 removed needless complication for modern SMT solvers
blanchet
parents: 58429
diff changeset
   322
    @{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
   323
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   324
  exception BAD_PATTERN of unit
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
  fun is_builtin_conn_or_pred ctxt c ts =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   327
    is_some (SMT_Builtin.dest_builtin_conn ctxt c ts) orelse
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   328
    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
   329
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   330
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   331
fun folify ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   332
  let
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   333
    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
   334
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   335
    fun in_term pat t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   336
      (case Term.strip_comb t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   337
        (@{const True}, []) => t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   338
      | (@{const False}, []) => t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   339
      | (u as Const (@{const_name If}, _), [t1, t2, t3]) =>
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56096
diff changeset
   340
          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
   341
      | (Const (c as (n, _)), ts) =>
61782
7d754aca6a75 removed needless complication for modern SMT solvers
blanchet
parents: 58429
diff changeset
   342
          if is_builtin_conn_or_pred ctxt c ts orelse is_quant n then
7d754aca6a75 removed needless complication for modern SMT solvers
blanchet
parents: 58429
diff changeset
   343
            if pat then raise BAD_PATTERN () else in_form t
7d754aca6a75 removed needless complication for modern SMT solvers
blanchet
parents: 58429
diff changeset
   344
          else
7d754aca6a75 removed needless complication for modern SMT solvers
blanchet
parents: 58429
diff changeset
   345
            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
   346
      | (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
   347
      | _ => t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   348
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   349
    and in_pat ((p as Const (@{const_name pat}, _)) $ t) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   350
          p $ in_term true t
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   351
      | in_pat ((p as Const (@{const_name nopat}, _)) $ t) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   352
          p $ in_term true t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   353
      | in_pat t = raise TERM ("bad pattern", [t])
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   354
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   355
    and in_pats ps =
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   356
      in_list @{typ "pattern symb_list"} (SOME o in_list @{typ pattern} (try in_pat)) ps
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   357
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   358
    and in_trigger ((c as @{const trigger}) $ p $ t) = c $ in_pats p $ in_form t
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
   359
      | in_trigger t = in_form t
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   360
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   361
    and in_form t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   362
      (case Term.strip_comb t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   363
        (q as Const (qn, _), [Abs (n, T, u)]) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   364
          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
   365
          else in_term false t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   366
      | (Const c, ts) =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   367
          (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
   368
            SOME (_, _, us, mk) => mk (map in_form us)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   369
          | NONE =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   370
              (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
   371
                SOME (_, _, us, mk) => mk (map (in_term false) us)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   372
              | NONE => in_term false t))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   373
      | _ => in_term false t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   374
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   375
    map in_form #>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   376
    pair (fol_rules, I)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   377
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   378
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   379
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   380
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   381
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   382
(* translation into intermediate format *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   383
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   384
(** utility functions **)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   385
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   386
val quantifier = (fn
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   387
    @{const_name All} => SOME SForall
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   388
  | @{const_name Ex} => SOME SExists
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   389
  | _ => NONE)
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
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
   392
      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
   393
  | group_quant _ Ts t = (Ts, t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   394
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   395
fun dest_pat (Const (@{const_name pat}, _) $ t) = (t, true)
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   396
  | dest_pat (Const (@{const_name nopat}, _) $ t) = (t, false)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   397
  | dest_pat t = raise TERM ("bad pattern", [t])
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   398
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   399
fun dest_pats [] = I
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   400
  | dest_pats ts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   401
      (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
   402
        (ps, [true]) => cons (SPat ps)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   403
      | (ps, [false]) => cons (SNoPat ps)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   404
      | _ => raise TERM ("bad multi-pattern", ts))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   405
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   406
fun dest_trigger (@{const trigger} $ tl $ t) =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   407
      (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
   408
  | dest_trigger t = ([], t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   409
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   410
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
   411
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   412
    val (Ts, u) = group_quant qn [T] t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   413
    val (ps, p) = dest_trigger u
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
   414
  in (q, rev Ts, ps, p) end)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   415
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   416
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
   417
  | 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
   418
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   419
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   420
(** translation from Isabelle terms into SMT intermediate terms **)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   421
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   422
fun intermediate logic dtyps builtin ctxt ts trx =
56078
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
    fun transT (T as TFree _) = add_typ T true
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   425
      | 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
   426
      | transT (T as Type _) =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   427
          (case SMT_Builtin.dest_builtin_typ ctxt T of
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   428
            SOME n => pair n
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   429
          | NONE => add_typ T true)
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
    fun app n ts = SApp (n, ts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   432
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   433
    fun trans t =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   434
      (case Term.strip_comb t of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   435
        (Const (qn, _), [Abs (_, T, t1)]) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   436
          (case dest_quant qn T t1 of
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
   437
            SOME (q, Ts, ps, b) =>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   438
              fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>>
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56811
diff changeset
   439
              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
   440
          | NONE => raise TERM ("unsupported quantifier", [t]))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   441
      | (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   442
          transT T ##>> trans t1 ##>> trans t2 #>> (fn ((U, u1), u2) => SLet (U, u1, u2))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   443
      | (u as Const (c as (_, T)), ts) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   444
          (case builtin ctxt c ts of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   445
            SOME (n, _, us, _) => fold_map trans us #>> app n
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   446
          | NONE => transs u T ts)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   447
      | (u as Free (_, T), ts) => transs u T ts
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   448
      | (Bound i, []) => pair (SVar i)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   449
      | _ => raise TERM ("bad SMT term", [t]))
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   450
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   451
    and transs t T ts =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   452
      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
   453
      in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   454
        fold_map transT Us ##>> transT U #-> (fn Up =>
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56096
diff changeset
   455
          add_fun t (SOME Up) ##>> fold_map trans ts #>> SApp)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   456
      end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   457
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   458
    val (us, trx') = fold_map trans ts trx
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   459
  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
   460
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   461
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   462
(* translation *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   463
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   464
structure Configs = Generic_Data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   465
(
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   466
  type T = (Proof.context -> config) SMT_Util.dict
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   467
  val empty = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   468
  val extend = I
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   469
  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
   470
)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   471
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   472
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
   473
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   474
fun get_config ctxt =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   475
  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
   476
  in
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   477
    (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
   478
      SOME cfg => cfg ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   479
    | NONE => error ("SMT: no translation configuration found " ^
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   480
        "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
   481
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   482
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57238
diff changeset
   483
fun translate ctxt smt_options comments ithms =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   484
  let
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
   485
    val {logic, fp_kinds, serialize} = get_config ctxt
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   486
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   487
    fun no_dtyps (tr_context, ctxt) ts =
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   488
      ((Termtab.empty, [], tr_context, ctxt), ts)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   489
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   490
    val ts1 = map (Envir.beta_eta_contract o SMT_Util.prop_of o snd) ithms
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   491
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   492
    val ((funcs, dtyps, tr_context, ctxt1), ts2) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   493
      ((empty_tr_context, ctxt), ts1)
58361
7f2b3b6f6ad1 added codatatype support for CVC4
blanchet
parents: 58360
diff changeset
   494
      |-> (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
   495
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   496
    fun is_binder (Const (@{const_name Let}, _) $ _) = true
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   497
      | is_binder t = Lambda_Lifting.is_quantifier t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   498
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   499
    fun mk_trigger ((q as Const (@{const_name All}, _)) $ Abs (n, T, t)) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   500
          q $ Abs (n, T, mk_trigger t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   501
      | mk_trigger (eq as (Const (@{const_name HOL.eq}, T) $ lhs $ _)) =
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   502
          Term.domain_type T --> @{typ pattern}
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   503
          |> (fn T => Const (@{const_name pat}, T) $ lhs)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   504
          |> SMT_Util.mk_symb_list @{typ pattern} o single
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57541
diff changeset
   505
          |> SMT_Util.mk_symb_list @{typ "pattern symb_list"} o single
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   506
          |> (fn t => @{const trigger} $ t $ eq)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   507
      | mk_trigger t = t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   508
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
   509
    val (ctxt2, (ts3, ll_defs)) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   510
      ts2
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   511
      |> eta_expand ctxt1 funcs
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   512
      |> rpair ctxt1
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   513
      |-> Lambda_Lifting.lift_lambdas NONE is_binder
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
   514
      |-> (fn (ts', ll_defs) => fn ctxt' =>
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
   515
          (ctxt', (intro_explicit_application ctxt' funcs (map mk_trigger ll_defs @ ts'), ll_defs)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   516
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   517
    val ((rewrite_rules, builtin), ts4) = folify ctxt2 ts3
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56096
diff changeset
   518
      |>> apfst (cons fun_app_eq)
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   519
val _ = 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
   520
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   521
    (ts4, tr_context)
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   522
    |-> intermediate logic dtyps (builtin SMT_Builtin.dest_builtin) ctxt2
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57238
diff changeset
   523
    |>> uncurry (serialize smt_options comments)
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 57540
diff changeset
   524
    ||> 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
   525
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   526
57229
blanchet
parents: 57165
diff changeset
   527
end;