src/HOL/Tools/Metis/metis_instantiations.ML
author wenzelm
Wed, 25 Jun 2025 16:35:25 +0200
changeset 82768 8f866fd6fae1
parent 82574 318526b74e6f
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Metis/metis_instantiations.ML
81757
4d15005da582 tuned documentation and order of instantiated facts
Lukas Bartl
parents: 81746
diff changeset
     2
    Author:     Lukas Bartl, Universitaet Augsburg
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     3
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     4
Inference of Variable Instantiations for Metis.
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     5
*)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     6
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     7
signature METIS_INSTANTIATIONS =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     8
sig
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
     9
  type inst
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    10
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    11
  type infer_params = {
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
    12
    infer_ctxt : Proof.context,
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    13
    type_enc : string,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    14
    lam_trans : string,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    15
    th_name : thm -> string option,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    16
    new_skolem : bool,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    17
    th_cls_pairs : (thm * thm list) list,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    18
    lifted : (string * term) list,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    19
    sym_tab : int Symtab.table,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    20
    axioms : (Metis_Thm.thm * thm) list,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    21
    mth : Metis_Thm.thm
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    22
  }
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    23
81746
8b4340d82248 Rename "suggest_of" to "instantiate"
Lukas Bartl
parents: 81254
diff changeset
    24
  val instantiate : bool Config.T
8b4340d82248 Rename "suggest_of" to "instantiate"
Lukas Bartl
parents: 81254
diff changeset
    25
  val instantiate_undefined : bool Config.T
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    26
  val metis_call : string -> string -> string
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    27
  val table_of_thm_inst : thm * inst -> cterm Vars.table
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    28
  val pretty_name_inst : Proof.context -> string * inst -> Pretty.T
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    29
  val string_of_name_inst : Proof.context -> string * inst -> string
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
    30
  val infer_thm_insts : Proof.context -> infer_params -> (thm * inst) list option
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
    31
  val instantiate_call : Proof.context -> infer_params -> thm -> unit
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    32
end;
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    33
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    34
structure Metis_Instantiations : METIS_INSTANTIATIONS =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    35
struct
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    36
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    37
open ATP_Problem_Generate
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    38
open ATP_Proof_Reconstruct
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    39
open Metis_Generate
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    40
open Metis_Reconstruct
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    41
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    42
(* Type of variable instantiations of a theorem.  This is a list of (certified)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    43
 * terms suitably ordered for an of-instantiation of the theorem. *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    44
type inst = cterm list
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    45
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    46
(* Params needed for inference of variable instantiations *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    47
type infer_params = {
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
    48
  infer_ctxt : Proof.context,
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    49
  type_enc : string,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    50
  lam_trans : string,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    51
  th_name : thm -> string option,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    52
  new_skolem : bool,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    53
  th_cls_pairs : (thm * thm list) list,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    54
  lifted : (string * term) list,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    55
  sym_tab : int Symtab.table,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    56
  axioms : (Metis_Thm.thm * thm) list,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    57
  mth : Metis_Thm.thm
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    58
}
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    59
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    60
(* Config option to enable suggestion of of-instantiations (e.g. "assms(1)[of x]") *)
81746
8b4340d82248 Rename "suggest_of" to "instantiate"
Lukas Bartl
parents: 81254
diff changeset
    61
val instantiate = Attrib.setup_config_bool @{binding "metis_instantiate"} (K false)
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    62
(* Config option to allow instantiation of variables with "undefined" *)
81746
8b4340d82248 Rename "suggest_of" to "instantiate"
Lukas Bartl
parents: 81254
diff changeset
    63
val instantiate_undefined = Attrib.setup_config_bool @{binding "metis_instantiate_undefined"} (K true)
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    64
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    65
fun metis_call type_enc lam_trans =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    66
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    67
    val type_enc =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    68
      case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases type_enc of
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    69
        [alias] => alias
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    70
      | _ => type_enc
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    71
    val opts =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    72
      [] |> lam_trans <> default_metis_lam_trans ? cons lam_trans
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    73
         |> type_enc <> partial_typesN ? cons type_enc
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    74
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    75
    metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")")
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    76
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    77
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    78
(* Variables of a theorem ordered the same way as in of-instantiations
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    79
 * (cf. Rule_Insts.of_rule) *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    80
fun of_vars_of_thm th = Vars.build (Vars.add_vars (Thm.full_prop_of th)) |> Vars.list_set
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    81
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    82
(* "_" terms are used for variables without instantiation *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    83
val is_dummy_cterm = Term.is_dummy_pattern o Thm.term_of
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    84
val thm_insts_trivial = forall (forall is_dummy_cterm o snd)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    85
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    86
fun table_of_thm_inst (th, cts) =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    87
  of_vars_of_thm th ~~ cts
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    88
  |> filter_out (is_dummy_cterm o snd)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    89
  |> Vars.make
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    90
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    91
fun open_attrib name =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    92
  case try (unsuffix "]") name of
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    93
    NONE => name ^ "["
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    94
  | SOME name' => name' ^ ", "
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    95
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    96
fun pretty_inst_cterm ctxt ct =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    97
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    98
    val pretty = Syntax.pretty_term ctxt (Thm.term_of ct)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
    99
    val is_dummy = ATP_Util.content_of_pretty pretty = "_"
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   100
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   101
    (* Do not quote single "_" *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   102
    pretty |> not is_dummy ? ATP_Util.pretty_maybe_quote ctxt
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   103
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   104
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   105
(* Pretty of-instantiation *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   106
fun pretty_name_inst ctxt (name, inst) =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   107
  case drop_suffix is_dummy_cterm inst of
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   108
    [] => Pretty.str name
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   109
  | cts =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   110
    map (pretty_inst_cterm ctxt) cts
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   111
    |> Pretty.breaks o cons (Pretty.str "of")
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   112
    |> Pretty.enclose (open_attrib name) "]"
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   113
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   114
val string_of_name_inst = Pretty.string_of oo pretty_name_inst
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   115
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   116
(* Infer Metis axioms with corresponding subtitutions *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   117
fun infer_metis_axiom_substs mth =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   118
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   119
    fun add_metis_axiom_substs msubst mth =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   120
      case Metis_Thm.inference mth of
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   121
        (Metis_Thm.Axiom, []) => cons (msubst, mth)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   122
      | (Metis_Thm.Metis_Subst, _) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   123
        (case Metis_Proof.thmToInference mth of
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   124
          Metis_Proof.Metis_Subst (msubst', par) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   125
            add_metis_axiom_substs (Metis_Subst.compose msubst' msubst) par
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   126
        | _ => raise Fail "expected Metis_Subst")
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   127
      | (_, pars) => fold (add_metis_axiom_substs msubst) pars
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   128
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   129
    add_metis_axiom_substs Metis_Subst.empty mth []
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   130
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   131
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   132
(* Variables are renamed during clausification by importing and exporting
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   133
 * theorems.  Do the same here to find the right variable. *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   134
fun import_export_thm ctxt th =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   135
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   136
    (* cf. Meson_Clausify.nnf_axiom *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   137
    val (import_th, import_ctxt) =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   138
      Drule.zero_var_indexes th
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   139
      |> (fn th' => Variable.import true [th'] ctxt)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   140
      |>> the_single o snd
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   141
    (* cf. Meson_Clausify.cnf_axiom *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   142
    val export_th = import_th
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   143
      |> singleton (Variable.export import_ctxt ctxt)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   144
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   145
    (* cf. Meson.finish_cnf *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   146
    Drule.zero_var_indexes export_th
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   147
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   148
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   149
(* Find the theorem corresponding to a Metis axiom if it has a name.
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   150
 * "Theorems" are Isabelle theorems given to the metis tactic.
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   151
 * "Axioms" are clauses given to the Metis prover. *)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   152
fun metis_axiom_to_thm ctxt {th_name, th_cls_pairs, axioms, ...} (msubst, maxiom) =
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   153
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   154
    val axiom = lookth axioms maxiom
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   155
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   156
    List.find (have_common_thm ctxt [axiom] o snd) th_cls_pairs
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   157
    |> Option.mapPartial (fn (th, _) => th_name th |> Option.map (pair th))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   158
    |> Option.map (pair (msubst, maxiom, axiom))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   159
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   160
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   161
(* Build a table : term Vartab.table list Thmtab.table that maps a theorem to a
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   162
 * list of variable substitutions (theorems can be instantiated multiple times)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   163
 * from Metis substitutions *)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   164
fun metis_substs_to_table ctxt {infer_ctxt, type_enc, lifted, new_skolem, sym_tab, ...}
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   165
    th_of_vars th_msubsts =
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   166
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   167
    val _ = trace_msg ctxt (K "AXIOM SUBSTITUTIONS")
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   168
    val type_enc = type_enc_of_string Strict type_enc
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   169
    val thy = Proof_Context.theory_of ctxt
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   170
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   171
    (* Replace Skolem terms with "_" because they are unknown to the user
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   172
     * (cf. Metis_Generate.reveal_old_skolem_terms) *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   173
    val dummy_old_skolem_terms = map_aterms
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   174
      (fn t as Const (s, T) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   175
          if String.isPrefix old_skolem_const_prefix s
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   176
          then Term.dummy_pattern T
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   177
          else t
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   178
        | t => t)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   179
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   180
    (* Infer types and replace "_" terms/types with schematic variables *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   181
    val infer_types_pattern =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   182
      singleton (Type_Infer_Context.infer_types_finished
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   183
        (Proof_Context.set_mode Proof_Context.mode_pattern infer_ctxt))
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   184
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   185
    (* "undefined" if allowed and not using new_skolem, "_" otherwise. *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   186
    val undefined_pattern =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   187
      (* new_skolem uses schematic variables which should not be instantiated with "undefined" *)
81746
8b4340d82248 Rename "suggest_of" to "instantiate"
Lukas Bartl
parents: 81254
diff changeset
   188
      if not new_skolem andalso Config.get ctxt instantiate_undefined then
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   189
        Const o (pair @{const_name "undefined"})
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   190
      else
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   191
        Term.dummy_pattern
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   192
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   193
    (* Instantiate schematic and free variables *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   194
    val inst_vars_frees = map_aterms
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   195
      (fn t as Free (x, T) =>
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   196
        (* an undeclared/internal free variable is unknown/inaccessible to the user,
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   197
         * so we replace it with "_" *)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   198
        let
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   199
          val x' = Variable.revert_fixed ctxt x
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   200
        in
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   201
          if not (Variable.is_declared ctxt x') orelse Name.is_internal x' then
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   202
            Term.dummy_pattern T
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   203
          else
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   204
            t
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   205
        end
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   206
      | Var (_, T) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   207
        (* a schematic variable can be replaced with any term,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   208
         * so we replace it with "undefined" *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   209
        undefined_pattern T
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   210
      | t => t)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   211
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   212
    (* Remove arguments of "_" unknown functions because the functions could
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   213
     * change (e.g. instantiations can change Skolem functions).
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   214
     * Type inference also introduces arguments (cf. Term.replace_dummy_patterns). *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   215
    fun eliminate_dummy_args (t $ u) =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   216
        let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   217
          val t' = eliminate_dummy_args t
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   218
        in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   219
          if Term.is_dummy_pattern t' then
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   220
            Term.dummy_pattern (Term.fastype_of t' |> Term.range_type)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   221
          else
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   222
            t' $ eliminate_dummy_args u
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   223
        end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   224
      | eliminate_dummy_args (Abs (s, T, t)) = Abs (s, T, eliminate_dummy_args t)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   225
      | eliminate_dummy_args t = t
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   226
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   227
    (* Polish Isabelle term, s.t. it can be displayed
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   228
     * (cf. Metis_Reconstruct.polish_hol_terms and ATP_Proof_Reconstruct.termify_line) *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   229
    val polish_term =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   230
      reveal_lam_lifted lifted (* reveal lifted lambdas *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   231
      #> dummy_old_skolem_terms (* eliminate Skolem terms *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   232
      #> infer_types_pattern (* type inference *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   233
      #> eliminate_lam_wrappers (* eliminate Metis.lambda wrappers *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   234
      #> uncombine_term thy (* eliminate Meson.COMB* terms *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   235
      #> Envir.beta_eta_contract (* simplify lambda terms *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   236
      #> simplify_bool (* simplify boolean terms *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   237
      #> Term.show_dummy_patterns (* reveal "_" that have been replaced by type inference *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   238
      #> inst_vars_frees (* eliminate schematic and free variables *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   239
      #> eliminate_dummy_args (* eliminate arguments of "_" (unknown functions) *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   240
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   241
    (* Translate a Metis substitution to Isabelle *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   242
    fun add_subst_to_table ((msubst, maxiom, axiom), (th, name)) th_substs_table =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   243
      let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   244
        val _ = trace_msg ctxt (fn () => "Metis axiom: " ^ Metis_Thm.toString maxiom)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   245
        val _ = trace_msg ctxt (fn () => "Metis substitution: " ^ Metis_Subst.toString msubst)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   246
        val _ = trace_msg ctxt (fn () => "Isabelle axiom: " ^ Thm.string_of_thm infer_ctxt axiom)
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   247
        val _ = trace_msg ctxt (fn () => "Isabelle theorem: " ^
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   248
          Thm.string_of_thm ctxt th ^ " (" ^ name ^ ")")
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   249
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   250
        (* Only indexnames of variables without types
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   251
         * because types can change during clausification *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   252
        val vars =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   253
          inter (op =) (map fst (of_vars_of_thm axiom))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   254
            (map fst (Thmtab.lookup th_of_vars th |> the))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   255
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   256
        (* Translate Metis variable and term to Isabelle *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   257
        fun translate_var_term (mvar, mt) =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   258
          Metis_Name.toString mvar
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   259
          (* cf. remove_typeinst in Metis_Reconstruct.inst_inference *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   260
          |> unprefix_and_unascii schematic_var_prefix
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   261
          (* cf. find_var in Metis_Reconstruct.inst_inference *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   262
          |> Option.mapPartial (fn name => List.find (fn (a, _) => a = name) vars)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   263
          (* cf. subst_translation in Metis_Reconstruct.inst_inference *)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   264
          |> Option.map (fn var => (var, hol_term_of_metis infer_ctxt type_enc sym_tab mt))
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   265
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   266
        (* Build the substitution table and instantiate the remaining schematic variables *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   267
        fun build_subst_table substs =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   268
          let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   269
            (* Variable of the axiom that didn't get instantiated by Metis,
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   270
             * the type is later set via type unification *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   271
            val undefined_term = undefined_pattern (TVar ((Name.aT, 0), []))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   272
          in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   273
            Vartab.build (vars |> fold (fn var => Vartab.default
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   274
              (var, AList.lookup (op =) substs var |> the_default undefined_term)))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   275
          end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   276
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   277
        val raw_substs =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   278
          Metis_Subst.toList msubst
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   279
          |> map_filter translate_var_term
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   280
        val _ = if null raw_substs then () else
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   281
          trace_msg ctxt (fn () => cat_lines ("Raw Isabelle substitution:" ::
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   282
            (raw_substs |> map (fn (var, t) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   283
              Term.string_of_vname var ^ " |-> " ^
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   284
              Syntax.string_of_term infer_ctxt t))))
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   285
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   286
        val subst = raw_substs
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   287
          |> map (apsnd polish_term)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   288
          |> build_subst_table
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   289
        val _ = trace_msg ctxt (fn () =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   290
          if Vartab.is_empty subst then
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   291
            "No Isabelle substitution"
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   292
          else
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   293
            cat_lines ("Final Isabelle substitution:" ::
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   294
              (Vartab.dest subst |> map (fn (var ,t) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   295
                Term.string_of_vname var ^ " |-> " ^
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   296
                Syntax.string_of_term ctxt t))))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   297
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   298
        (* Try to merge the substitution with another substitution of the theorem *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   299
        fun insert_subst [] = [subst]
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   300
          | insert_subst (subst' :: substs') =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   301
            Vartab.merge Term.aconv_untyped (subst', subst) :: substs'
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   302
            handle Vartab.DUP _ => subst' :: insert_subst substs'
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   303
      in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   304
        Thmtab.lookup th_substs_table th
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   305
        |> insert_subst o these
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   306
        |> (fn substs => Thmtab.update (th, substs) th_substs_table)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   307
      end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   308
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   309
    fold add_subst_to_table th_msubsts Thmtab.empty
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   310
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   311
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   312
(* Build variable instantiations : thm * inst list from the table *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   313
fun table_to_thm_insts ctxt th_of_vars th_substs_table =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   314
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   315
    val thy = Proof_Context.theory_of ctxt
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   316
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   317
    (* Unify types of a variable with the term for instantiation
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   318
     * (cf. Drule.infer_instantiate_types, Metis_Reconstruct.cterm_incr_types) *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   319
    fun unify (tyenv, maxidx) T t =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   320
      let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   321
        val t' = Term.map_types (Logic.incr_tvar (maxidx + 1)) t
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   322
        val U = Term.fastype_of t'
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   323
        val maxidx' = Term.maxidx_term t' maxidx
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   324
      in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   325
        (t', Sign.typ_unify thy (T, U) (tyenv, maxidx'))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   326
      end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   327
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   328
    (* Instantiate type variables with a type unifier and import remaining ones
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   329
     * (cf. Drule.infer_instantiate_types) *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   330
    fun inst_unifier (ts, tyenv) =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   331
      let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   332
        val instT =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   333
          TVars.build (tyenv |> Vartab.fold (fn (xi, (S, T)) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   334
            TVars.add ((xi, S), Envir.norm_type tyenv T)))
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   335
        val ts' = map (Term_Subst.instantiate (instT, Vars.empty)) ts
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   336
      in
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   337
        Variable.importT_terms ts' ctxt |> fst
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   338
      end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   339
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   340
    (* Build variable instantiations from a substitution table and instantiate
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   341
     * the remaining schematic variables with "_" *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   342
    fun add_subst th subst =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   343
      let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   344
        val of_vars = Thmtab.lookup th_of_vars th |> the
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   345
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   346
        fun unify_or_dummy (var, T) unify_params =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   347
          Vartab.lookup subst var
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   348
          |> Option.map (unify unify_params T)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   349
          |> the_default (Term.dummy_pattern T, unify_params)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   350
      in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   351
        fold_map unify_or_dummy of_vars (Vartab.empty, Thm.maxidx_of th)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   352
        |> inst_unifier o apsnd fst
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   353
        |> map (Thm.cterm_of ctxt)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   354
        |> cons o pair th
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   355
      end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   356
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   357
    Thmtab.fold_rev (fn (th, substs) => fold (add_subst th) substs) th_substs_table []
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   358
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   359
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   360
fun really_infer_thm_insts ctxt (infer_params as {th_name, th_cls_pairs, mth, ...}) =
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   361
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   362
    (* Compute the theorem variables ordered as in of-instantiations.
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   363
     * import_export_thm is used to get the same variables as in axioms. *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   364
    val th_of_vars =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   365
      Thmtab.build (th_cls_pairs |> fold (fn (th, _) => Thmtab.default
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   366
        (th, of_vars_of_thm (import_export_thm ctxt th))))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   367
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   368
    val th_insts =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   369
      infer_metis_axiom_substs mth
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   370
      |> map_filter (metis_axiom_to_thm ctxt infer_params)
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   371
      |> metis_substs_to_table ctxt infer_params th_of_vars
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   372
      |> table_to_thm_insts ctxt th_of_vars
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   373
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   374
    val _ = trace_msg ctxt (fn () => cat_lines ("THEOREM INSTANTIATIONS" ::
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   375
     (if thm_insts_trivial th_insts then
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   376
        ["No instantiation available"]
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   377
      else
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   378
        th_insts |> maps (fn th_inst as (th, _) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   379
          let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   380
            val table = table_of_thm_inst th_inst
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   381
          in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   382
            if Vars.is_empty table then
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   383
              []
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   384
            else
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   385
              "Theorem " ^ Thm.string_of_thm ctxt th ^ " (" ^ the (th_name th) ^ "):" ::
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   386
              (Vars.dest table |> map (fn (var, ct) =>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   387
                Syntax.string_of_term ctxt (Var var) ^ " |-> " ^
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   388
                Syntax.string_of_term ctxt (Thm.term_of ct)))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   389
          end))))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   390
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   391
    th_insts
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   392
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   393
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   394
(* Infer variable instantiations *)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   395
fun infer_thm_insts ctxt infer_params =
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   396
  \<^try>\<open>SOME (really_infer_thm_insts ctxt infer_params)
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   397
    catch exn => (trace_msg ctxt (fn () => Runtime.exn_message exn); NONE)\<close>
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   398
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   399
(* Write suggested command with of-instantiations *)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   400
fun write_command ctxt {type_enc, lam_trans, th_name, ...} st th_insts =
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   401
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   402
    val _ = trace_msg ctxt (K "SUGGEST OF-INSTANTIATIONS")
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   403
    val apply_str = if Thm.nprems_of st = 0 then "by" else "apply"
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   404
    val command_str = th_insts
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   405
      |> map (pretty_name_inst ctxt o apfst (the o th_name))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   406
      |> cons (Pretty.str (metis_call type_enc lam_trans))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   407
      |> Pretty.enclose (apply_str ^ " (") ")" o Pretty.breaks
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   408
      |> Pretty.symbolic_string_of (* markup string *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   409
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   410
    writeln ("Found variable instantiations, try this:" ^
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   411
      (* Add optional markup break (command may need multiple lines) *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   412
      Markup.markup (Markup.break {width = 1, indent = 2}) " " ^
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   413
      Active.sendback_markup_command command_str)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   414
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   415
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   416
(* Infer variable instantiations and suggest of-instantiations *)
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   417
fun instantiate_call ctxt infer_params st =
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   418
  infer_thm_insts ctxt infer_params
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   419
  |> Option.mapPartial (Option.filter (not o thm_insts_trivial))
82574
318526b74e6f tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents: 81757
diff changeset
   420
  |> Option.app (write_command ctxt infer_params st)
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   421
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents:
diff changeset
   422
end;