src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
author blanchet
Tue, 06 Nov 2012 14:46:21 +0100
changeset 51033 4ea26c74d7ea
parent 51032 d9c1b11a78d2
child 51034 930a10e674ef
permissions -rw-r--r--
use implications rather than disjunctions to improve readability
blanchet@50898
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
blanchet@50898
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@50898
     3
    Author:     Steffen Juilf Smolka, TU Muenchen
blanchet@50898
     4
blanchet@50929
     5
Isar proof reconstruction from ATP proofs.
blanchet@50898
     6
*)
blanchet@50898
     7
blanchet@50898
     8
signature SLEDGEHAMMER_PROOF_RECONSTRUCT =
blanchet@50898
     9
sig
blanchet@50929
    10
  type 'a proof = 'a ATP_Proof.proof
blanchet@50929
    11
  type stature = ATP_Problem_Generate.stature
blanchet@50929
    12
blanchet@50929
    13
  datatype reconstructor =
blanchet@50929
    14
    Metis of string * string |
blanchet@50929
    15
    SMT
blanchet@50929
    16
blanchet@50929
    17
  datatype play =
blanchet@50929
    18
    Played of reconstructor * Time.time |
blanchet@50929
    19
    Trust_Playable of reconstructor * Time.time option |
blanchet@50929
    20
    Failed_to_Play of reconstructor
blanchet@50929
    21
blanchet@50929
    22
  type minimize_command = string list -> string
blanchet@50929
    23
  type one_line_params =
blanchet@50929
    24
    play * string * (string * stature) list * minimize_command * int * int
blanchet@50929
    25
  type isar_params =
blanchet@51019
    26
    bool * bool * Time.time * real * string Symtab.table
blanchet@51019
    27
    * (string * stature) list vector * int Symtab.table * string proof * thm
blanchet@50929
    28
blanchet@50929
    29
  val smtN : string
blanchet@50929
    30
  val string_for_reconstructor : reconstructor -> string
blanchet@50931
    31
  val thms_of_name : Proof.context -> string -> thm list
blanchet@50929
    32
  val lam_trans_from_atp_proof : string proof -> string -> string
blanchet@50929
    33
  val is_typed_helper_used_in_atp_proof : string proof -> bool
blanchet@50929
    34
  val used_facts_in_atp_proof :
blanchet@50929
    35
    Proof.context -> (string * stature) list vector -> string proof ->
blanchet@50929
    36
    (string * stature) list
blanchet@50929
    37
  val used_facts_in_unsound_atp_proof :
blanchet@50929
    38
    Proof.context -> (string * stature) list vector -> 'a proof ->
blanchet@50929
    39
    string list option
blanchet@50929
    40
  val one_line_proof_text : int -> one_line_params -> string
blanchet@50898
    41
  val isar_proof_text :
blanchet@50928
    42
    Proof.context -> bool -> isar_params -> one_line_params -> string
blanchet@50898
    43
  val proof_text :
blanchet@50928
    44
    Proof.context -> bool -> isar_params -> int -> one_line_params -> string
blanchet@50898
    45
end;
blanchet@50898
    46
blanchet@50898
    47
structure Sledgehammer_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT =
blanchet@50898
    48
struct
blanchet@50898
    49
blanchet@50898
    50
open ATP_Util
blanchet@50929
    51
open ATP_Problem
blanchet@50898
    52
open ATP_Proof
blanchet@50898
    53
open ATP_Problem_Generate
blanchet@50898
    54
open ATP_Proof_Reconstruct
blanchet@50933
    55
open Sledgehammer_Util
blanchet@50929
    56
blanchet@50929
    57
structure String_Redirect = ATP_Proof_Redirect(
blanchet@50929
    58
  type key = step_name
blanchet@50929
    59
  val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s')
blanchet@50929
    60
  val string_of = fst)
blanchet@50929
    61
blanchet@50898
    62
open String_Redirect
blanchet@50898
    63
blanchet@50931
    64
blanchet@50929
    65
(** reconstructors **)
blanchet@50929
    66
blanchet@50929
    67
datatype reconstructor =
blanchet@50929
    68
  Metis of string * string |
blanchet@50929
    69
  SMT
blanchet@50929
    70
blanchet@50929
    71
datatype play =
blanchet@50929
    72
  Played of reconstructor * Time.time |
blanchet@50929
    73
  Trust_Playable of reconstructor * Time.time option |
blanchet@50929
    74
  Failed_to_Play of reconstructor
blanchet@50929
    75
blanchet@50929
    76
val smtN = "smt"
blanchet@50929
    77
blanchet@50929
    78
fun string_for_reconstructor (Metis (type_enc, lam_trans)) =
blanchet@50929
    79
    metis_call type_enc lam_trans
blanchet@50929
    80
  | string_for_reconstructor SMT = smtN
blanchet@50929
    81
blanchet@50931
    82
fun thms_of_name ctxt name =
blanchet@50931
    83
  let
blanchet@50931
    84
    val lex = Keyword.get_lexicons
blanchet@50931
    85
    val get = maps (Proof_Context.get_fact ctxt o fst)
blanchet@50931
    86
  in
blanchet@50931
    87
    Source.of_string name
blanchet@50931
    88
    |> Symbol.source
blanchet@50996
    89
    |> Token.source {do_recover = SOME false} lex Position.start
blanchet@50931
    90
    |> Token.source_proper
blanchet@50931
    91
    |> Source.source Token.stopper (Parse_Spec.xthms1 >> get) NONE
blanchet@50931
    92
    |> Source.exhaust
blanchet@50931
    93
  end
blanchet@50931
    94
blanchet@50929
    95
blanchet@50929
    96
(** fact extraction from ATP proofs **)
blanchet@50929
    97
blanchet@50929
    98
fun find_first_in_list_vector vec key =
blanchet@50929
    99
  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
blanchet@50929
   100
                 | (_, value) => value) NONE vec
blanchet@50929
   101
blanchet@50929
   102
val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
blanchet@50929
   103
blanchet@50929
   104
fun resolve_one_named_fact fact_names s =
blanchet@50929
   105
  case try (unprefix fact_prefix) s of
blanchet@50929
   106
    SOME s' =>
blanchet@50929
   107
    let val s' = s' |> unprefix_fact_number |> unascii_of in
blanchet@50929
   108
      s' |> find_first_in_list_vector fact_names |> Option.map (pair s')
blanchet@50929
   109
    end
blanchet@50929
   110
  | NONE => NONE
blanchet@50929
   111
fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names)
blanchet@50929
   112
fun is_fact fact_names = not o null o resolve_fact fact_names
blanchet@50929
   113
blanchet@50929
   114
fun resolve_one_named_conjecture s =
blanchet@50929
   115
  case try (unprefix conjecture_prefix) s of
blanchet@50929
   116
    SOME s' => Int.fromString s'
blanchet@50929
   117
  | NONE => NONE
blanchet@50929
   118
blanchet@50929
   119
val resolve_conjecture = map_filter resolve_one_named_conjecture
blanchet@50929
   120
val is_conjecture = not o null o resolve_conjecture
blanchet@50929
   121
blanchet@50929
   122
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix
blanchet@50929
   123
blanchet@50929
   124
(* overapproximation (good enough) *)
blanchet@50929
   125
fun is_lam_lifted s =
blanchet@50929
   126
  String.isPrefix fact_prefix s andalso
blanchet@50929
   127
  String.isSubstring ascii_of_lam_fact_prefix s
blanchet@50929
   128
blanchet@50929
   129
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix)
blanchet@50929
   130
blanchet@50929
   131
fun is_axiom_used_in_proof pred =
blanchet@51027
   132
  exists (fn Inference_Step ((_, ss), _, _, _, []) => exists pred ss
blanchet@51027
   133
           | _ => false)
blanchet@50929
   134
blanchet@50929
   135
fun lam_trans_from_atp_proof atp_proof default =
blanchet@50929
   136
  case (is_axiom_used_in_proof is_combinator_def atp_proof,
blanchet@50929
   137
        is_axiom_used_in_proof is_lam_lifted atp_proof) of
blanchet@50929
   138
    (false, false) => default
blanchet@50929
   139
  | (false, true) => liftingN
blanchet@50929
   140
(*  | (true, true) => combs_and_liftingN -- not supported by "metis" *)
blanchet@50929
   141
  | (true, _) => combsN
blanchet@50929
   142
blanchet@50929
   143
val is_typed_helper_name =
blanchet@50929
   144
  String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
blanchet@50929
   145
fun is_typed_helper_used_in_atp_proof atp_proof =
blanchet@50929
   146
  is_axiom_used_in_proof is_typed_helper_name atp_proof
blanchet@50929
   147
blanchet@50929
   148
fun add_non_rec_defs fact_names accum =
blanchet@50929
   149
  Vector.foldl (fn (facts, facts') =>
blanchet@50929
   150
      union (op =) (filter (fn (_, (_, status)) => status = Non_Rec_Def) facts)
blanchet@50929
   151
            facts')
blanchet@50929
   152
    accum fact_names
blanchet@50929
   153
blanchet@50929
   154
val isa_ext = Thm.get_name_hint @{thm ext}
blanchet@50929
   155
val isa_short_ext = Long_Name.base_name isa_ext
blanchet@50929
   156
blanchet@50929
   157
fun ext_name ctxt =
blanchet@50929
   158
  if Thm.eq_thm_prop (@{thm ext},
blanchet@50929
   159
         singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
blanchet@50929
   160
    isa_short_ext
blanchet@50929
   161
  else
blanchet@50929
   162
    isa_ext
blanchet@50929
   163
blanchet@50929
   164
val leo2_ext = "extcnf_equal_neg"
blanchet@50929
   165
val leo2_unfold_def = "unfold_def"
blanchet@50929
   166
blanchet@51027
   167
fun add_fact ctxt fact_names (Inference_Step ((_, ss), _, _, rule, deps)) =
blanchet@50929
   168
    (if rule = leo2_ext then
blanchet@50929
   169
       insert (op =) (ext_name ctxt, (Global, General))
blanchet@50929
   170
     else if rule = leo2_unfold_def then
blanchet@50929
   171
       (* LEO 1.3.3 does not record definitions properly, leading to missing
blanchet@50929
   172
         dependencies in the TSTP proof. Remove the next line once this is
blanchet@50929
   173
         fixed. *)
blanchet@50929
   174
       add_non_rec_defs fact_names
blanchet@50929
   175
     else if rule = satallax_coreN then
blanchet@50929
   176
       (fn [] =>
blanchet@50929
   177
           (* Satallax doesn't include definitions in its unsatisfiable cores,
blanchet@50929
   178
              so we assume the worst and include them all here. *)
blanchet@50929
   179
           [(ext_name ctxt, (Global, General))] |> add_non_rec_defs fact_names
blanchet@50929
   180
         | facts => facts)
blanchet@50929
   181
     else
blanchet@50929
   182
       I)
blanchet@50929
   183
    #> (if null deps then union (op =) (resolve_fact fact_names ss)
blanchet@50929
   184
        else I)
blanchet@50929
   185
  | add_fact _ _ _ = I
blanchet@50929
   186
blanchet@50929
   187
fun used_facts_in_atp_proof ctxt fact_names atp_proof =
blanchet@50929
   188
  if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names
blanchet@50929
   189
  else fold (add_fact ctxt fact_names) atp_proof []
blanchet@50929
   190
blanchet@50929
   191
fun used_facts_in_unsound_atp_proof _ _ [] = NONE
blanchet@50929
   192
  | used_facts_in_unsound_atp_proof ctxt fact_names atp_proof =
blanchet@50929
   193
    let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in
blanchet@50929
   194
      if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso
blanchet@50929
   195
         not (is_axiom_used_in_proof (is_conjecture o single) atp_proof) then
blanchet@50929
   196
        SOME (map fst used_facts)
blanchet@50929
   197
      else
blanchet@50929
   198
        NONE
blanchet@50929
   199
    end
blanchet@50929
   200
blanchet@50929
   201
blanchet@50929
   202
(** one-liner reconstructor proofs **)
blanchet@50929
   203
blanchet@50929
   204
fun string_for_label (s, num) = s ^ string_of_int num
blanchet@50929
   205
blanchet@50929
   206
fun show_time NONE = ""
blanchet@50929
   207
  | show_time (SOME ext_time) = " (" ^ string_from_ext_time ext_time ^ ")"
blanchet@50929
   208
blanchet@51001
   209
(* FIXME: Various bugs, esp. with "unfolding"
blanchet@50929
   210
fun unusing_chained_facts _ 0 = ""
blanchet@50929
   211
  | unusing_chained_facts used_chaineds num_chained =
blanchet@50929
   212
    if length used_chaineds = num_chained then ""
blanchet@50929
   213
    else if null used_chaineds then "(* using no facts *) "
blanchet@50929
   214
    else "(* using only " ^ space_implode " " used_chaineds ^ " *) "
blanchet@51001
   215
*)
blanchet@50929
   216
blanchet@50929
   217
fun apply_on_subgoal _ 1 = "by "
blanchet@50929
   218
  | apply_on_subgoal 1 _ = "apply "
blanchet@50929
   219
  | apply_on_subgoal i n =
blanchet@50929
   220
    "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n
blanchet@50929
   221
blanchet@50929
   222
fun using_labels [] = ""
blanchet@50929
   223
  | using_labels ls =
blanchet@50929
   224
    "using " ^ space_implode " " (map string_for_label ls) ^ " "
blanchet@50929
   225
blanchet@50929
   226
fun command_call name [] =
blanchet@50929
   227
    name |> not (Lexicon.is_identifier name) ? enclose "(" ")"
blanchet@50929
   228
  | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
blanchet@50929
   229
blanchet@50929
   230
fun reconstructor_command reconstr i n used_chaineds num_chained (ls, ss) =
blanchet@51001
   231
  (* unusing_chained_facts used_chaineds num_chained ^ *)
blanchet@50929
   232
  using_labels ls ^ apply_on_subgoal i n ^
blanchet@50929
   233
  command_call (string_for_reconstructor reconstr) ss
blanchet@50929
   234
blanchet@50929
   235
fun try_command_line banner time command =
blanchet@50929
   236
  banner ^ ": " ^ Markup.markup Isabelle_Markup.sendback command ^
blanchet@50929
   237
  show_time time ^ "."
blanchet@50929
   238
blanchet@50929
   239
fun minimize_line _ [] = ""
blanchet@50929
   240
  | minimize_line minimize_command ss =
blanchet@50929
   241
    case minimize_command ss of
blanchet@50929
   242
      "" => ""
blanchet@50929
   243
    | command =>
blanchet@50929
   244
      "\nTo minimize: " ^ Markup.markup Isabelle_Markup.sendback command ^ "."
blanchet@50929
   245
blanchet@50929
   246
fun split_used_facts facts =
blanchet@50929
   247
  facts |> List.partition (fn (_, (sc, _)) => sc = Chained)
blanchet@50929
   248
        |> pairself (sort_distinct (string_ord o pairself fst))
blanchet@50929
   249
blanchet@50929
   250
type minimize_command = string list -> string
blanchet@50929
   251
type one_line_params =
blanchet@50929
   252
  play * string * (string * stature) list * minimize_command * int * int
blanchet@50929
   253
blanchet@50929
   254
fun one_line_proof_text num_chained
blanchet@50929
   255
        (preplay, banner, used_facts, minimize_command, subgoal,
blanchet@50929
   256
         subgoal_count) =
blanchet@50929
   257
  let
blanchet@50929
   258
    val (chained, extra) = split_used_facts used_facts
blanchet@50929
   259
    val (failed, reconstr, ext_time) =
blanchet@50929
   260
      case preplay of
blanchet@50929
   261
        Played (reconstr, time) => (false, reconstr, (SOME (false, time)))
blanchet@50929
   262
      | Trust_Playable (reconstr, time) =>
blanchet@50929
   263
        (false, reconstr,
blanchet@50929
   264
         case time of
blanchet@50929
   265
           NONE => NONE
blanchet@50929
   266
         | SOME time =>
blanchet@50929
   267
           if time = Time.zeroTime then NONE else SOME (true, time))
blanchet@50929
   268
      | Failed_to_Play reconstr => (true, reconstr, NONE)
blanchet@50929
   269
    val try_line =
blanchet@50929
   270
      ([], map fst extra)
blanchet@50929
   271
      |> reconstructor_command reconstr subgoal subgoal_count (map fst chained)
blanchet@50929
   272
                               num_chained
blanchet@50929
   273
      |> (if failed then
blanchet@50929
   274
            enclose "One-line proof reconstruction failed: "
blanchet@50929
   275
                     ".\n(Invoking \"sledgehammer\" with \"[strict]\" might \
blanchet@50929
   276
                     \solve this.)"
blanchet@50929
   277
          else
blanchet@50929
   278
            try_command_line banner ext_time)
blanchet@50929
   279
  in try_line ^ minimize_line minimize_command (map fst (extra @ chained)) end
blanchet@50929
   280
blanchet@50929
   281
blanchet@50929
   282
(** Isar proof construction and manipulation **)
blanchet@50929
   283
blanchet@50929
   284
type label = string * int
blanchet@50929
   285
type facts = label list * string list
blanchet@50929
   286
blanchet@50929
   287
datatype isar_qualifier = Show | Then | Moreover | Ultimately
blanchet@50929
   288
blanchet@50929
   289
datatype isar_step =
blanchet@50929
   290
  Fix of (string * typ) list |
blanchet@50929
   291
  Let of term * term |
blanchet@50929
   292
  Assume of label * term |
blanchet@50929
   293
  Prove of isar_qualifier list * label * term * byline
blanchet@50929
   294
and byline =
blanchet@50929
   295
  By_Metis of facts |
blanchet@50929
   296
  Case_Split of isar_step list list * facts
blanchet@50929
   297
blanchet@51032
   298
val assume_prefix = "a"
blanchet@50929
   299
val have_prefix = "f"
blanchet@50929
   300
val raw_prefix = "x"
blanchet@50929
   301
blanchet@50929
   302
fun raw_label_for_name (num, ss) =
blanchet@50929
   303
  case resolve_conjecture ss of
blanchet@50929
   304
    [j] => (conjecture_prefix, j)
blanchet@50929
   305
  | _ => (raw_prefix ^ ascii_of num, 0)
blanchet@50929
   306
blanchet@51020
   307
fun label_of_clause [name] = raw_label_for_name name
blanchet@51020
   308
  | label_of_clause c = (space_implode "___" (map fst c), 0)
blanchet@51020
   309
blanchet@51020
   310
fun add_fact_from_dependencies fact_names (names as [(_, ss)]) =
blanchet@51020
   311
    if is_fact fact_names ss then
blanchet@51020
   312
      apsnd (union (op =) (map fst (resolve_fact fact_names ss)))
blanchet@51020
   313
    else
blanchet@51020
   314
      apfst (insert (op =) (label_of_clause names))
blanchet@51020
   315
  | add_fact_from_dependencies fact_names names =
blanchet@51020
   316
    apfst (insert (op =) (label_of_clause names))
blanchet@50929
   317
blanchet@50929
   318
fun repair_name "$true" = "c_True"
blanchet@50929
   319
  | repair_name "$false" = "c_False"
blanchet@50929
   320
  | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
blanchet@50929
   321
  | repair_name s =
blanchet@50929
   322
    if is_tptp_equal s orelse
blanchet@50929
   323
       (* seen in Vampire proofs *)
blanchet@50929
   324
       (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
blanchet@50929
   325
      tptp_equal
blanchet@50929
   326
    else
blanchet@50929
   327
      s
blanchet@50929
   328
blanchet@50929
   329
fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
blanchet@50929
   330
  | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
blanchet@50929
   331
blanchet@50929
   332
fun infer_formula_types ctxt =
blanchet@50929
   333
  Type.constraint HOLogic.boolT
blanchet@50929
   334
  #> Syntax.check_term
blanchet@50929
   335
         (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
blanchet@50929
   336
blanchet@50929
   337
val combinator_table =
blanchet@50929
   338
  [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
blanchet@50929
   339
   (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
blanchet@50929
   340
   (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
blanchet@50929
   341
   (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
blanchet@50929
   342
   (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
blanchet@50929
   343
blanchet@50929
   344
fun uncombine_term thy =
blanchet@50929
   345
  let
blanchet@50929
   346
    fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
blanchet@50929
   347
      | aux (Abs (s, T, t')) = Abs (s, T, aux t')
blanchet@50929
   348
      | aux (t as Const (x as (s, _))) =
blanchet@50929
   349
        (case AList.lookup (op =) combinator_table s of
blanchet@50929
   350
           SOME thm => thm |> prop_of |> specialize_type thy x
blanchet@50929
   351
                           |> Logic.dest_equals |> snd
blanchet@50929
   352
         | NONE => t)
blanchet@50929
   353
      | aux t = t
blanchet@50929
   354
  in aux end
blanchet@50929
   355
blanchet@50929
   356
fun decode_line sym_tab (Definition_Step (name, phi1, phi2)) ctxt =
blanchet@50929
   357
    let
blanchet@50929
   358
      val thy = Proof_Context.theory_of ctxt
blanchet@50929
   359
      val t1 = prop_from_atp ctxt true sym_tab phi1
blanchet@50929
   360
      val vars = snd (strip_comb t1)
blanchet@50929
   361
      val frees = map unvarify_term vars
blanchet@50929
   362
      val unvarify_args = subst_atomic (vars ~~ frees)
blanchet@50929
   363
      val t2 = prop_from_atp ctxt true sym_tab phi2
blanchet@50929
   364
      val (t1, t2) =
blanchet@50929
   365
        HOLogic.eq_const HOLogic.typeT $ t1 $ t2
blanchet@50929
   366
        |> unvarify_args |> uncombine_term thy |> infer_formula_types ctxt
blanchet@50929
   367
        |> HOLogic.dest_eq
blanchet@50929
   368
    in
blanchet@50929
   369
      (Definition_Step (name, t1, t2),
blanchet@50929
   370
       fold Variable.declare_term (maps Misc_Legacy.term_frees [t1, t2]) ctxt)
blanchet@50929
   371
    end
blanchet@51027
   372
  | decode_line sym_tab (Inference_Step (name, role, u, rule, deps)) ctxt =
blanchet@50929
   373
    let
blanchet@50929
   374
      val thy = Proof_Context.theory_of ctxt
blanchet@50929
   375
      val t = u |> prop_from_atp ctxt true sym_tab
blanchet@50929
   376
                |> uncombine_term thy |> infer_formula_types ctxt
blanchet@50929
   377
    in
blanchet@51027
   378
      (Inference_Step (name, role, t, rule, deps),
blanchet@50929
   379
       fold Variable.declare_term (Misc_Legacy.term_frees t) ctxt)
blanchet@50929
   380
    end
blanchet@50929
   381
fun decode_lines ctxt sym_tab lines =
blanchet@50929
   382
  fst (fold_map (decode_line sym_tab) lines ctxt)
blanchet@50929
   383
blanchet@50929
   384
fun replace_one_dependency (old, new) dep =
blanchet@50929
   385
  if is_same_atp_step dep old then new else [dep]
blanchet@50929
   386
fun replace_dependencies_in_line _ (line as Definition_Step _) = line
blanchet@51027
   387
  | replace_dependencies_in_line p
blanchet@51027
   388
        (Inference_Step (name, role, t, rule, deps)) =
blanchet@51027
   389
    Inference_Step (name, role, t, rule,
blanchet@50929
   390
                    fold (union (op =) o replace_one_dependency p) deps [])
blanchet@50929
   391
blanchet@50929
   392
(* No "real" literals means only type information (tfree_tcs, clsrel, or
blanchet@50929
   393
   clsarity). *)
blanchet@50929
   394
fun is_only_type_information t = t aconv @{term True}
blanchet@50929
   395
blanchet@50929
   396
fun is_same_inference _ (Definition_Step _) = false
blanchet@51027
   397
  | is_same_inference t (Inference_Step (_, _, t', _, _)) = t aconv t'
blanchet@50929
   398
blanchet@50929
   399
(* Discard facts; consolidate adjacent lines that prove the same formula, since
blanchet@50929
   400
   they differ only in type information.*)
blanchet@50929
   401
fun add_line _ (line as Definition_Step _) lines = line :: lines
blanchet@51027
   402
  | add_line fact_names (Inference_Step (name as (_, ss), role, t, rule, []))
blanchet@51027
   403
             lines =
blanchet@50929
   404
    (* No dependencies: fact, conjecture, or (for Vampire) internal facts or
blanchet@50929
   405
       definitions. *)
blanchet@50929
   406
    if is_fact fact_names ss then
blanchet@50929
   407
      (* Facts are not proof lines. *)
blanchet@50929
   408
      if is_only_type_information t then
blanchet@50929
   409
        map (replace_dependencies_in_line (name, [])) lines
blanchet@50929
   410
      (* Is there a repetition? If so, replace later line by earlier one. *)
blanchet@50929
   411
      else case take_prefix (not o is_same_inference t) lines of
blanchet@50929
   412
        (_, []) => lines (* no repetition of proof line *)
blanchet@51027
   413
      | (pre, Inference_Step (name', _, _, _, _) :: post) =>
blanchet@50929
   414
        pre @ map (replace_dependencies_in_line (name', [name])) post
blanchet@50929
   415
      | _ => raise Fail "unexpected inference"
blanchet@50929
   416
    else if is_conjecture ss then
blanchet@51027
   417
      Inference_Step (name, role, t, rule, []) :: lines
blanchet@50929
   418
    else
blanchet@50929
   419
      map (replace_dependencies_in_line (name, [])) lines
blanchet@51027
   420
  | add_line _ (Inference_Step (name, role, t, rule, deps)) lines =
blanchet@50929
   421
    (* Type information will be deleted later; skip repetition test. *)
blanchet@50929
   422
    if is_only_type_information t then
blanchet@51027
   423
      Inference_Step (name, role, t, rule, deps) :: lines
blanchet@50929
   424
    (* Is there a repetition? If so, replace later line by earlier one. *)
blanchet@50929
   425
    else case take_prefix (not o is_same_inference t) lines of
blanchet@50929
   426
      (* FIXME: Doesn't this code risk conflating proofs involving different
blanchet@50929
   427
         types? *)
blanchet@51027
   428
       (_, []) => Inference_Step (name, role, t, rule, deps) :: lines
blanchet@51027
   429
     | (pre, Inference_Step (name', role, t', rule, _) :: post) =>
blanchet@51027
   430
       Inference_Step (name, role, t', rule, deps) ::
blanchet@50929
   431
       pre @ map (replace_dependencies_in_line (name', [name])) post
blanchet@50929
   432
     | _ => raise Fail "unexpected inference"
blanchet@50929
   433
blanchet@50929
   434
val waldmeister_conjecture_num = "1.0.0.0"
blanchet@50929
   435
blanchet@50929
   436
val repair_waldmeister_endgame =
blanchet@50929
   437
  let
blanchet@51027
   438
    fun do_tail (Inference_Step (name, role, t, rule, deps)) =
blanchet@51027
   439
        Inference_Step (name, role, s_not t, rule, deps)
blanchet@50929
   440
      | do_tail line = line
blanchet@50929
   441
    fun do_body [] = []
blanchet@51027
   442
      | do_body ((line as Inference_Step ((num, _), _, _, _, _)) :: lines) =
blanchet@50929
   443
        if num = waldmeister_conjecture_num then map do_tail (line :: lines)
blanchet@50929
   444
        else line :: do_body lines
blanchet@50929
   445
      | do_body (line :: lines) = line :: do_body lines
blanchet@50929
   446
  in do_body end
blanchet@50929
   447
blanchet@50929
   448
(* Recursively delete empty lines (type information) from the proof. *)
blanchet@51027
   449
fun add_nontrivial_line (line as Inference_Step (name, _, t, _, [])) lines =
blanchet@50929
   450
    if is_only_type_information t then delete_dependency name lines
blanchet@50929
   451
    else line :: lines
blanchet@50929
   452
  | add_nontrivial_line line lines = line :: lines
blanchet@50929
   453
and delete_dependency name lines =
blanchet@50929
   454
  fold_rev add_nontrivial_line
blanchet@50929
   455
           (map (replace_dependencies_in_line (name, [])) lines) []
blanchet@50929
   456
blanchet@50929
   457
(* ATPs sometimes reuse free variable names in the strangest ways. Removing
blanchet@50929
   458
   offending lines often does the trick. *)
blanchet@50929
   459
fun is_bad_free frees (Free x) = not (member (op =) frees x)
blanchet@50929
   460
  | is_bad_free _ _ = false
blanchet@50929
   461
blanchet@50933
   462
fun add_desired_line _ _ (line as Definition_Step (name, _, _)) (j, lines) =
blanchet@50929
   463
    (j, line :: map (replace_dependencies_in_line (name, [])) lines)
blanchet@50933
   464
  | add_desired_line fact_names frees
blanchet@51027
   465
        (Inference_Step (name as (_, ss), role, t, rule, deps)) (j, lines) =
blanchet@50929
   466
    (j + 1,
blanchet@50929
   467
     if is_fact fact_names ss orelse
blanchet@50929
   468
        is_conjecture ss orelse
blanchet@50929
   469
        (* the last line must be kept *)
blanchet@50929
   470
        j = 0 orelse
blanchet@50929
   471
        (not (is_only_type_information t) andalso
blanchet@50929
   472
         null (Term.add_tvars t []) andalso
blanchet@50929
   473
         not (exists_subterm (is_bad_free frees) t) andalso
blanchet@50933
   474
         length deps >= 2 andalso
blanchet@50929
   475
         (* kill next to last line, which usually results in a trivial step *)
blanchet@50929
   476
         j <> 1) then
blanchet@51027
   477
       Inference_Step (name, role, t, rule, deps) :: lines  (* keep line *)
blanchet@50929
   478
     else
blanchet@50929
   479
       map (replace_dependencies_in_line (name, deps)) lines)  (* drop line *)
blanchet@50929
   480
blanchet@50898
   481
(** Type annotations **)
blanchet@50898
   482
blanchet@50898
   483
fun post_traverse_term_type' f _ (t as Const (_, T)) s = f t T s
blanchet@50898
   484
  | post_traverse_term_type' f _ (t as Free (_, T)) s = f t T s
blanchet@50898
   485
  | post_traverse_term_type' f _ (t as Var (_, T)) s = f t T s
blanchet@50898
   486
  | post_traverse_term_type' f env (t as Bound i) s = f t (nth env i) s
blanchet@50898
   487
  | post_traverse_term_type' f env (Abs (x, T1, b)) s =
blanchet@50898
   488
    let
blanchet@50898
   489
      val ((b', s'), T2) = post_traverse_term_type' f (T1 :: env) b s
blanchet@50898
   490
    in f (Abs (x, T1, b')) (T1 --> T2) s' end
blanchet@50898
   491
  | post_traverse_term_type' f env (u $ v) s =
blanchet@50898
   492
    let
blanchet@50898
   493
      val ((u', s'), Type (_, [_, T])) = post_traverse_term_type' f env u s
blanchet@50898
   494
      val ((v', s''), _) = post_traverse_term_type' f env v s'
blanchet@50898
   495
    in f (u' $ v') T s'' end
blanchet@50898
   496
blanchet@50898
   497
fun post_traverse_term_type f s t =
blanchet@50898
   498
  post_traverse_term_type' (fn t => fn T => fn s => (f t T s, T)) [] t s |> fst
blanchet@50898
   499
fun post_fold_term_type f s t =
blanchet@50898
   500
  post_traverse_term_type (fn t => fn T => fn s => (t, f t T s)) s t |> snd
blanchet@50898
   501
blanchet@50898
   502
(* Data structures, orders *)
blanchet@50898
   503
val cost_ord = prod_ord int_ord (prod_ord int_ord int_ord)
blanchet@50898
   504
blanchet@50898
   505
structure Var_Set_Tab = Table(
blanchet@50898
   506
  type key = indexname list
blanchet@50898
   507
  val ord = list_ord Term_Ord.fast_indexname_ord)
blanchet@50898
   508
blanchet@50898
   509
(* (1) Generalize Types *)
blanchet@50898
   510
fun generalize_types ctxt t =
blanchet@50898
   511
  t |> map_types (fn _ => dummyT)
blanchet@50898
   512
    |> Syntax.check_term
blanchet@50898
   513
         (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)
blanchet@50898
   514
blanchet@50898
   515
(* (2) Typing-spot Table *)
blanchet@50898
   516
local
blanchet@51019
   517
fun key_of_atype (TVar (z, _)) =
blanchet@51019
   518
    Ord_List.insert Term_Ord.fast_indexname_ord z
blanchet@50898
   519
  | key_of_atype _ = I
blanchet@50898
   520
fun key_of_type T = fold_atyps key_of_atype T []
blanchet@50898
   521
fun update_tab t T (tab, pos) =
blanchet@50898
   522
  (case key_of_type T of
blanchet@50898
   523
     [] => tab
blanchet@50898
   524
   | key =>
blanchet@50898
   525
     let val cost = (size_of_typ T, (size_of_term t, pos)) in
blanchet@50898
   526
       case Var_Set_Tab.lookup tab key of
blanchet@50898
   527
         NONE => Var_Set_Tab.update_new (key, cost) tab
blanchet@50898
   528
       | SOME old_cost =>
blanchet@50898
   529
         (case cost_ord (cost, old_cost) of
blanchet@50898
   530
            LESS => Var_Set_Tab.update (key, cost) tab
blanchet@50898
   531
          | _ => tab)
blanchet@50898
   532
     end,
blanchet@50898
   533
   pos + 1)
blanchet@50898
   534
in
blanchet@50898
   535
val typing_spot_table =
blanchet@50898
   536
  post_fold_term_type update_tab (Var_Set_Tab.empty, 0) #> fst
blanchet@50898
   537
end
blanchet@50898
   538
blanchet@50898
   539
(* (3) Reverse-Greedy *)
blanchet@50898
   540
fun reverse_greedy typing_spot_tab =
blanchet@50898
   541
  let
blanchet@50898
   542
    fun update_count z =
blanchet@50898
   543
      fold (fn tvar => fn tab =>
blanchet@50898
   544
        let val c = Vartab.lookup tab tvar |> the_default 0 in
blanchet@50898
   545
          Vartab.update (tvar, c + z) tab
blanchet@50898
   546
        end)
blanchet@50898
   547
    fun superfluous tcount =
blanchet@50898
   548
      forall (fn tvar => the (Vartab.lookup tcount tvar) > 1)
blanchet@50898
   549
    fun drop_superfluous (tvars, (_, (_, spot))) (spots, tcount) =
blanchet@50898
   550
      if superfluous tcount tvars then (spots, update_count ~1 tvars tcount)
blanchet@50898
   551
      else (spot :: spots, tcount)
blanchet@50898
   552
    val (typing_spots, tvar_count_tab) =
blanchet@50898
   553
      Var_Set_Tab.fold
blanchet@50898
   554
        (fn kv as (k, _) => apfst (cons kv) #> apsnd (update_count 1 k))
blanchet@50898
   555
        typing_spot_tab ([], Vartab.empty)
blanchet@50898
   556
      |>> sort_distinct (rev_order o cost_ord o pairself snd)
blanchet@50898
   557
  in fold drop_superfluous typing_spots ([], tvar_count_tab) |> fst end
blanchet@50898
   558
blanchet@50898
   559
(* (4) Introduce Annotations *)
blanchet@50898
   560
fun introduce_annotations thy spots t t' =
blanchet@50898
   561
  let
blanchet@50898
   562
    val get_types = post_fold_term_type (K cons) []
blanchet@50898
   563
    fun match_types tp =
blanchet@50898
   564
      fold (Sign.typ_match thy) (op ~~ (pairself get_types tp)) Vartab.empty
blanchet@50898
   565
    fun unica' b x [] = if b then [x] else []
blanchet@50898
   566
      | unica' b x (y :: ys) =
blanchet@50898
   567
        if x = y then unica' false x ys
blanchet@50898
   568
        else unica' true y ys |> b ? cons x
blanchet@50898
   569
    fun unica ord xs =
blanchet@50898
   570
      case sort ord xs of x :: ys => unica' true x ys | [] => []
blanchet@50898
   571
    val add_all_tfree_namesT = fold_atyps (fn TFree (x, _) => cons x | _ => I)
blanchet@50898
   572
    fun erase_unica_tfrees env =
blanchet@50898
   573
      let
blanchet@50898
   574
        val unica =
blanchet@50898
   575
          Vartab.fold (add_all_tfree_namesT o snd o snd) env []
blanchet@50898
   576
          |> unica fast_string_ord
blanchet@50898
   577
        val erase_unica = map_atyps
blanchet@50898
   578
          (fn T as TFree (s, _) =>
blanchet@50898
   579
              if Ord_List.member fast_string_ord unica s then dummyT else T
blanchet@50898
   580
            | T => T)
blanchet@50898
   581
      in Vartab.map (K (apsnd erase_unica)) env end
blanchet@50898
   582
    val env = match_types (t', t) |> erase_unica_tfrees
blanchet@50898
   583
    fun get_annot env (TFree _) = (false, (env, dummyT))
blanchet@50898
   584
      | get_annot env (T as TVar (v, S)) =
blanchet@50898
   585
        let val T' = Envir.subst_type env T in
blanchet@50898
   586
          if T' = dummyT then (false, (env, dummyT))
blanchet@50898
   587
          else (true, (Vartab.update (v, (S, dummyT)) env, T'))
blanchet@50898
   588
        end
blanchet@50898
   589
      | get_annot env (Type (S, Ts)) =
blanchet@50898
   590
        (case fold_rev (fn T => fn (b, (env, Ts)) =>
blanchet@50898
   591
                  let
blanchet@50898
   592
                    val (b', (env', T)) = get_annot env T
blanchet@50898
   593
                  in (b orelse b', (env', T :: Ts)) end)
blanchet@50898
   594
                Ts (false, (env, [])) of
blanchet@50898
   595
           (true, (env', Ts)) => (true, (env', Type (S, Ts)))
blanchet@50898
   596
         | (false, (env', _)) => (false, (env', dummyT)))
blanchet@50898
   597
    fun post1 _ T (env, cp, ps as p :: ps', annots) =
blanchet@50898
   598
        if p <> cp then
blanchet@50898
   599
          (env, cp + 1, ps, annots)
blanchet@50898
   600
        else
blanchet@50898
   601
          let val (_, (env', T')) = get_annot env T in
blanchet@50898
   602
            (env', cp + 1, ps', (p, T') :: annots)
blanchet@50898
   603
          end
blanchet@50898
   604
      | post1 _ _ accum = accum
blanchet@50898
   605
    val (_, _, _, annots) = post_fold_term_type post1 (env, 0, spots, []) t'
blanchet@50898
   606
    fun post2 t _ (cp, annots as (p, T) :: annots') =
blanchet@50898
   607
        if p <> cp then (t, (cp + 1, annots))
blanchet@50898
   608
        else (Type.constraint T t, (cp + 1, annots'))
blanchet@50898
   609
      | post2 t _ x = (t, x)
blanchet@50898
   610
  in post_traverse_term_type post2 (0, rev annots) t |> fst end
blanchet@50898
   611
blanchet@50898
   612
(* (5) Annotate *)
blanchet@50898
   613
fun annotate_types ctxt t =
blanchet@50898
   614
  let
blanchet@50898
   615
    val thy = Proof_Context.theory_of ctxt
blanchet@50898
   616
    val t' = generalize_types ctxt t
blanchet@50898
   617
    val typing_spots =
blanchet@50898
   618
      t' |> typing_spot_table
blanchet@50898
   619
         |> reverse_greedy
blanchet@50898
   620
         |> sort int_ord
blanchet@50898
   621
  in introduce_annotations thy typing_spots t t' end
blanchet@50898
   622
blanchet@50929
   623
val indent_size = 2
blanchet@50929
   624
val no_label = ("", ~1)
blanchet@50929
   625
blanchet@50898
   626
fun string_for_proof ctxt type_enc lam_trans i n =
blanchet@50898
   627
  let
blanchet@50898
   628
    fun fix_print_mode f x =
blanchet@50898
   629
      Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
blanchet@50898
   630
                               (print_mode_value ())) f x
blanchet@50898
   631
    fun do_indent ind = replicate_string (ind * indent_size) " "
blanchet@50898
   632
    fun do_free (s, T) =
blanchet@50898
   633
      maybe_quote s ^ " :: " ^
blanchet@50898
   634
      maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
blanchet@50898
   635
    fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
blanchet@50898
   636
    fun do_have qs =
blanchet@50898
   637
      (if member (op =) qs Moreover then "moreover " else "") ^
blanchet@50898
   638
      (if member (op =) qs Ultimately then "ultimately " else "") ^
blanchet@50898
   639
      (if member (op =) qs Then then
blanchet@50898
   640
         if member (op =) qs Show then "thus" else "hence"
blanchet@50898
   641
       else
blanchet@50898
   642
         if member (op =) qs Show then "show" else "have")
blanchet@50898
   643
    val do_term =
blanchet@50898
   644
      maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
blanchet@50898
   645
      o annotate_types ctxt
blanchet@50898
   646
    val reconstr = Metis (type_enc, lam_trans)
blanchet@50936
   647
    fun do_facts ind (ls, ss) =
blanchet@50936
   648
      "\n" ^ do_indent (ind + 1) ^
blanchet@50898
   649
      reconstructor_command reconstr 1 1 [] 0
blanchet@50898
   650
          (ls |> sort_distinct (prod_ord string_ord int_ord),
blanchet@50898
   651
           ss |> sort_distinct string_ord)
blanchet@50898
   652
    and do_step ind (Fix xs) =
blanchet@50898
   653
        do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
blanchet@50898
   654
      | do_step ind (Let (t1, t2)) =
blanchet@50898
   655
        do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
blanchet@50898
   656
      | do_step ind (Assume (l, t)) =
blanchet@50898
   657
        do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
blanchet@50898
   658
      | do_step ind (Prove (qs, l, t, By_Metis facts)) =
blanchet@50898
   659
        do_indent ind ^ do_have qs ^ " " ^
blanchet@50936
   660
        do_label l ^ do_term t ^ do_facts ind facts ^ "\n"
blanchet@50898
   661
      | do_step ind (Prove (qs, l, t, Case_Split (proofs, facts))) =
blanchet@50898
   662
        implode (map (prefix (do_indent ind ^ "moreover\n") o do_block ind)
blanchet@50898
   663
                     proofs) ^
blanchet@50936
   664
        do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^
blanchet@50936
   665
        do_facts ind facts ^ "\n"
blanchet@50898
   666
    and do_steps prefix suffix ind steps =
blanchet@50898
   667
      let val s = implode (map (do_step ind) steps) in
blanchet@50898
   668
        replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
blanchet@50898
   669
        String.extract (s, ind * indent_size,
blanchet@50898
   670
                        SOME (size s - ind * indent_size - 1)) ^
blanchet@50898
   671
        suffix ^ "\n"
blanchet@50898
   672
      end
blanchet@50898
   673
    and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
blanchet@50898
   674
    (* One-step proofs are pointless; better use the Metis one-liner
blanchet@50898
   675
       directly. *)
blanchet@50898
   676
    and do_proof [Prove (_, _, _, By_Metis _)] = ""
blanchet@50898
   677
      | do_proof proof =
blanchet@50898
   678
        (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
blanchet@50898
   679
        do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^
blanchet@50898
   680
        (if n <> 1 then "next" else "qed")
blanchet@50898
   681
  in do_proof end
blanchet@50898
   682
blanchet@50929
   683
fun used_labels_of_step (Prove (_, _, _, by)) =
blanchet@50929
   684
    (case by of
blanchet@50929
   685
       By_Metis (ls, _) => ls
blanchet@50929
   686
     | Case_Split (proofs, (ls, _)) =>
blanchet@50929
   687
       fold (union (op =) o used_labels_of) proofs ls)
blanchet@50929
   688
  | used_labels_of_step _ = []
blanchet@50929
   689
and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
blanchet@50929
   690
blanchet@50929
   691
fun kill_useless_labels_in_proof proof =
blanchet@50929
   692
  let
blanchet@50929
   693
    val used_ls = used_labels_of proof
blanchet@50929
   694
    fun do_label l = if member (op =) used_ls l then l else no_label
blanchet@50929
   695
    fun do_step (Assume (l, t)) = Assume (do_label l, t)
blanchet@50929
   696
      | do_step (Prove (qs, l, t, by)) =
blanchet@50929
   697
        Prove (qs, do_label l, t,
blanchet@50929
   698
               case by of
blanchet@50929
   699
                 Case_Split (proofs, facts) =>
blanchet@50929
   700
                 Case_Split (map (map do_step) proofs, facts)
blanchet@50929
   701
               | _ => by)
blanchet@50929
   702
      | do_step step = step
blanchet@50929
   703
  in map do_step proof end
blanchet@50929
   704
blanchet@50929
   705
fun prefix_for_depth n = replicate_string (n + 1)
blanchet@50929
   706
blanchet@50929
   707
val relabel_proof =
blanchet@50929
   708
  let
blanchet@50929
   709
    fun aux _ _ _ [] = []
blanchet@51032
   710
      | aux subst depth (next_assum, next_have) (Assume (l, t) :: proof) =
blanchet@50929
   711
        if l = no_label then
blanchet@51032
   712
          Assume (l, t) :: aux subst depth (next_assum, next_have) proof
blanchet@50929
   713
        else
blanchet@51032
   714
          let val l' = (prefix_for_depth depth assume_prefix, next_assum) in
blanchet@50929
   715
            Assume (l', t) ::
blanchet@51032
   716
            aux ((l, l') :: subst) depth (next_assum + 1, next_have) proof
blanchet@50929
   717
          end
blanchet@51032
   718
      | aux subst depth (next_assum, next_have)
blanchet@50929
   719
            (Prove (qs, l, t, by) :: proof) =
blanchet@50929
   720
        let
blanchet@51032
   721
          val (l', subst, next_have) =
blanchet@50929
   722
            if l = no_label then
blanchet@51032
   723
              (l, subst, next_have)
blanchet@50929
   724
            else
blanchet@51032
   725
              let val l' = (prefix_for_depth depth have_prefix, next_have) in
blanchet@51032
   726
                (l', (l, l') :: subst, next_have + 1)
blanchet@51032
   727
              end
blanchet@50929
   728
          val relabel_facts =
blanchet@50929
   729
            apfst (maps (the_list o AList.lookup (op =) subst))
blanchet@50929
   730
          val by =
blanchet@50929
   731
            case by of
blanchet@50929
   732
              By_Metis facts => By_Metis (relabel_facts facts)
blanchet@50929
   733
            | Case_Split (proofs, facts) =>
blanchet@50929
   734
              Case_Split (map (aux subst (depth + 1) (1, 1)) proofs,
blanchet@50929
   735
                          relabel_facts facts)
blanchet@50929
   736
        in
blanchet@51032
   737
          Prove (qs, l', t, by) :: aux subst depth (next_assum, next_have) proof
blanchet@50929
   738
        end
blanchet@50929
   739
      | aux subst depth nextp (step :: proof) =
blanchet@50929
   740
        step :: aux subst depth nextp proof
blanchet@50929
   741
  in aux [] 0 (1, 1) end
blanchet@50929
   742
blanchet@50932
   743
val merge_timeout_slack = 1.2
blanchet@50932
   744
blanchet@51019
   745
val label_ord = prod_ord int_ord fast_string_ord o pairself swap
blanchet@51019
   746
blanchet@51019
   747
structure Label_Table = Table(
blanchet@51019
   748
  type key = label
blanchet@51019
   749
  val ord = label_ord)
blanchet@51019
   750
blanchet@51019
   751
fun shrink_proof debug ctxt type_enc lam_trans preplay
blanchet@51019
   752
                 preplay_timeout isar_shrinkage proof =
blanchet@50898
   753
  let
blanchet@51019
   754
    (* clean vector interface *)
blanchet@51019
   755
    fun get i v = Vector.sub (v, i)
blanchet@51019
   756
    fun replace x i v = Vector.update (v, i, x)
blanchet@51019
   757
    fun update f i v = replace (get i v |> f) i v
blanchet@51019
   758
    fun v_fold_index f v s =
blanchet@51019
   759
      Vector.foldl (fn (x, (i, s)) => (i+1, f (i, x) s)) (0, s) v |> snd
blanchet@51019
   760
blanchet@51019
   761
    (* Queue interface to table *)
blanchet@51019
   762
    fun pop tab key =
blanchet@51019
   763
      let val v = hd (Inttab.lookup_list tab key) in
blanchet@51019
   764
        (v, Inttab.remove_list (op =) (key, v) tab)
blanchet@51019
   765
      end
blanchet@51019
   766
    fun pop_max tab = pop tab (the (Inttab.max_key tab))
blanchet@51019
   767
    val is_empty = Inttab.is_empty
blanchet@51019
   768
    fun add_list tab xs = fold (Inttab.insert_list (op =)) xs tab
blanchet@51019
   769
blanchet@51019
   770
    (* proof vector *)
blanchet@51019
   771
    val proof_vect = proof |> map SOME |> Vector.fromList
blanchet@51019
   772
    val n = Vector.length proof_vect
blanchet@51019
   773
    val n_target = Real.fromInt n / isar_shrinkage |> Real.round
blanchet@51019
   774
blanchet@51019
   775
    (* table for mapping from label to proof position *)
blanchet@51019
   776
    fun update_table (i, Prove (_, label, _, _)) =
blanchet@51019
   777
        Label_Table.update_new (label, i)
blanchet@51019
   778
      | update_table _ = I
blanchet@51019
   779
    val label_index_table = fold_index update_table proof Label_Table.empty
blanchet@51019
   780
blanchet@51019
   781
    (* proof references *)
blanchet@51019
   782
    fun refs (Prove (_, _, _, By_Metis (refs, _))) =
blanchet@51019
   783
      map (the o Label_Table.lookup label_index_table) refs
blanchet@51019
   784
      | refs _ = []
blanchet@51019
   785
    val refed_by_vect =
blanchet@51019
   786
      Vector.tabulate (n, (fn _ => []))
blanchet@51019
   787
      |> fold_index (fn (i, step) => fold (update (cons i)) (refs step)) proof
blanchet@51019
   788
      |> Vector.map rev (* after rev, indices are sorted in ascending order *)
blanchet@51019
   789
blanchet@51019
   790
    (* candidates for elimination, use table as priority queue (greedy
blanchet@51019
   791
       algorithm) *)
blanchet@50898
   792
    fun cost (Prove (_, _ , t, _)) = Term.size_of_term t
blanchet@51019
   793
      | cost _ = 0
blanchet@51019
   794
    val cand_ord =  rev_order o prod_ord int_ord int_ord
blanchet@51019
   795
    val cand_tab =
blanchet@51019
   796
      v_fold_index
blanchet@51019
   797
        (fn (i, [_]) => cons (get i proof_vect |> the |> cost, i)
blanchet@51019
   798
        | _ => I) refed_by_vect []
blanchet@51019
   799
      |> Inttab.make_list
blanchet@50898
   800
blanchet@50932
   801
    (* Enrich context with local facts *)
blanchet@50898
   802
    val thy = Proof_Context.theory_of ctxt
blanchet@51019
   803
    fun enrich_ctxt' (Prove (_, label, t, _)) ctxt =
blanchet@51029
   804
        Proof_Context.put_thms false
blanchet@51029
   805
            (string_for_label label, SOME [Skip_Proof.make_thm thy t]) ctxt
blanchet@50898
   806
      | enrich_ctxt' _ ctxt = ctxt
blanchet@50898
   807
    val rich_ctxt = fold enrich_ctxt' proof ctxt
blanchet@50898
   808
blanchet@50898
   809
    (* Timing *)
blanchet@51019
   810
    fun take_time timeout tac arg =
blanchet@50932
   811
      let val timing = Timing.start () in
blanchet@51019
   812
        (TimeLimit.timeLimit timeout tac arg;
blanchet@51019
   813
         Timing.result timing |> #cpu |> SOME)
blanchet@51019
   814
        handle _ => NONE
blanchet@50898
   815
      end
blanchet@51019
   816
    val sum_up_time =
blanchet@51019
   817
      Vector.foldl
blanchet@51030
   818
        ((fn (SOME t, (b, s)) => (b, t + s)
blanchet@51030
   819
           | (NONE, (_, s)) => (true, preplay_timeout + s)) o apfst Lazy.force)
blanchet@51030
   820
        (false, seconds 0.0)
blanchet@51019
   821
blanchet@51019
   822
    (* Metis Preplaying *)
blanchet@51019
   823
    fun try_metis timeout (Prove (_, _, t, By_Metis fact_names)) =
blanchet@51019
   824
      if not preplay then (fn () => SOME (seconds 0.0)) else
blanchet@51019
   825
        let
blanchet@51019
   826
          val facts =
blanchet@51019
   827
            fact_names
blanchet@51019
   828
            |>> map string_for_label |> op @
blanchet@51019
   829
            |> map (the_single o thms_of_name rich_ctxt)
blanchet@51029
   830
          val goal =
blanchet@51029
   831
            Goal.prove (Config.put Metis_Tactic.verbose debug ctxt) [] [] t
blanchet@51019
   832
          fun tac {context = ctxt, prems = _} =
blanchet@51019
   833
            Metis_Tactic.metis_tac [type_enc] lam_trans ctxt facts 1
blanchet@51019
   834
        in
blanchet@51019
   835
          take_time timeout (fn () => goal tac)
blanchet@51019
   836
        end
blanchet@51019
   837
blanchet@51019
   838
    (* Lazy metis time vector, cache *)
blanchet@51019
   839
    val metis_time =
blanchet@51019
   840
      Vector.map (Lazy.lazy o try_metis preplay_timeout o the) proof_vect
blanchet@51019
   841
blanchet@51019
   842
    (* Merging *)
blanchet@51019
   843
    fun merge (Prove (qs1, label1, _, By_Metis (lfs1, gfs1)))
blanchet@51019
   844
              (Prove (qs2, label2 , t, By_Metis (lfs2, gfs2))) =
blanchet@50898
   845
      let
blanchet@51019
   846
        val qs = inter (op =) qs1 qs2 (* FIXME: Is this correct? *)
blanchet@50928
   847
          |> member (op =) (union (op =) qs1 qs2) Ultimately ? cons Ultimately
blanchet@50928
   848
          |> member (op =) qs2 Show ? cons Show
blanchet@51019
   849
        val ls = remove (op =) label1 lfs2 |> union (op =) lfs1
blanchet@51019
   850
        val ss = union (op =) gfs1 gfs2
blanchet@51019
   851
      in Prove (qs, label2, t, By_Metis (ls, ss)) end
blanchet@51019
   852
    fun try_merge metis_time (s1, i) (s2, j) =
blanchet@51019
   853
      (case get i metis_time |> Lazy.force of
blanchet@51019
   854
        NONE => (NONE, metis_time)
blanchet@51019
   855
      | SOME t1 =>
blanchet@51019
   856
        (case get j metis_time |> Lazy.force of
blanchet@51019
   857
          NONE => (NONE, metis_time)
blanchet@51019
   858
        | SOME t2 =>
blanchet@51019
   859
          let
blanchet@51019
   860
            val s12 = merge s1 s2
blanchet@51019
   861
            val timeout =
blanchet@51019
   862
              t1 + t2 |> Time.toReal |> curry Real.* merge_timeout_slack
blanchet@51019
   863
                      |> Time.fromReal
blanchet@51019
   864
          in
blanchet@51019
   865
            case try_metis timeout s12 () of
blanchet@51019
   866
              NONE => (NONE, metis_time)
blanchet@51019
   867
            | some_t12 =>
blanchet@51019
   868
              (SOME s12, metis_time
blanchet@51019
   869
                         |> replace (seconds 0.0 |> SOME |> Lazy.value) i
blanchet@51019
   870
                         |> replace (Lazy.value some_t12) j)
blanchet@51019
   871
blanchet@51019
   872
          end))
blanchet@51019
   873
blanchet@51019
   874
    fun merge_steps metis_time proof_vect refed_by cand_tab n' =
blanchet@51019
   875
      if is_empty cand_tab orelse n' <= n_target orelse n'<3 then
blanchet@51019
   876
        (sum_up_time metis_time,
blanchet@51019
   877
         Vector.foldr
blanchet@51019
   878
           (fn (NONE, proof) => proof | (SOME s, proof) => s :: proof)
blanchet@51019
   879
           [] proof_vect)
blanchet@51019
   880
      else
blanchet@51019
   881
        let
blanchet@51019
   882
          val (i, cand_tab) = pop_max cand_tab
blanchet@51019
   883
          val j = get i refed_by |> the_single
blanchet@51019
   884
          val s1 = get i proof_vect |> the
blanchet@51019
   885
          val s2 = get j proof_vect |> the
blanchet@51019
   886
        in
blanchet@51019
   887
          case try_merge metis_time (s1, i) (s2, j) of
blanchet@51019
   888
            (NONE, metis_time) =>
blanchet@51019
   889
            merge_steps metis_time proof_vect refed_by cand_tab n'
blanchet@51019
   890
          | (s, metis_time) => let
blanchet@51019
   891
            val refs = refs s1
blanchet@51019
   892
            val refed_by = refed_by |> fold
blanchet@51019
   893
              (update (Ord_List.remove int_ord i #> Ord_List.insert int_ord j)) refs
blanchet@51019
   894
            val new_candidates =
blanchet@51030
   895
              fold (fn (i, [_]) => cons (cost (get i proof_vect |> the), i)
blanchet@51030
   896
                     | _ => I)
blanchet@51019
   897
                (map (fn i => (i, get i refed_by)) refs) []
blanchet@51019
   898
            val cand_tab = add_list cand_tab new_candidates
blanchet@51019
   899
            val proof_vect = proof_vect |> replace NONE i |> replace s j
blanchet@51019
   900
          in
blanchet@51019
   901
            merge_steps metis_time proof_vect refed_by cand_tab (n' - 1)
blanchet@51019
   902
          end
blanchet@51019
   903
        end
blanchet@51019
   904
  in
blanchet@51019
   905
    merge_steps metis_time proof_vect refed_by_vect cand_tab n
blanchet@51019
   906
  end
blanchet@51019
   907
blanchet@51019
   908
val chain_direct_proof =
blanchet@51019
   909
  let
blanchet@51019
   910
    fun succedent_of_step (Prove (_, label, _, _)) = SOME label
blanchet@51019
   911
      | succedent_of_step (Assume (label, _)) = SOME label
blanchet@51019
   912
      | succedent_of_step _ = NONE
blanchet@51019
   913
    fun chain_inf (SOME label0)
blanchet@51019
   914
                  (step as Prove (qs, label, t, By_Metis (lfs, gfs))) =
blanchet@51019
   915
        if member (op =) lfs label0 then
blanchet@51019
   916
          Prove (Then :: qs, label, t,
blanchet@51019
   917
                 By_Metis (filter_out (curry (op =) label0) lfs, gfs))
blanchet@51019
   918
        else
blanchet@51019
   919
          step
blanchet@51019
   920
      | chain_inf _ (Prove (qs, label, t, Case_Split (proofs, facts))) =
blanchet@51019
   921
        Prove (qs, label, t, Case_Split ((map (chain_proof NONE) proofs), facts))
blanchet@51019
   922
      | chain_inf _ step = step
blanchet@51019
   923
    and chain_proof _ [] = []
blanchet@51019
   924
      | chain_proof (prev as SOME _) (i :: is) =
blanchet@51019
   925
        chain_inf prev i :: chain_proof (succedent_of_step i) is
blanchet@51019
   926
      | chain_proof _ (i :: is) =
blanchet@51019
   927
        i :: chain_proof (succedent_of_step i) is
blanchet@51019
   928
  in chain_proof NONE end
blanchet@50898
   929
blanchet@50929
   930
type isar_params =
blanchet@51019
   931
  bool * bool * Time.time * real * string Symtab.table
blanchet@51019
   932
  * (string * stature) list vector * int Symtab.table * string proof * thm
blanchet@50929
   933
blanchet@50933
   934
fun isar_proof_text ctxt isar_proofs
blanchet@51019
   935
    (debug, verbose, preplay_timeout, isar_shrinkage,
blanchet@51019
   936
     pool, fact_names, sym_tab, atp_proof, goal)
blanchet@50933
   937
    (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
blanchet@50898
   938
  let
blanchet@50898
   939
    val (params, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal
blanchet@50898
   940
    val frees = fold Term.add_frees (concl_t :: hyp_ts) []
blanchet@50898
   941
    val one_line_proof = one_line_proof_text 0 one_line_params
blanchet@50898
   942
    val type_enc =
blanchet@50898
   943
      if is_typed_helper_used_in_atp_proof atp_proof then full_typesN
blanchet@50898
   944
      else partial_typesN
blanchet@50898
   945
    val lam_trans = lam_trans_from_atp_proof atp_proof metis_default_lam_trans
blanchet@51019
   946
    val preplay = preplay_timeout <> seconds 0.0
blanchet@50898
   947
blanchet@50898
   948
    fun isar_proof_of () =
blanchet@50898
   949
      let
blanchet@50898
   950
        val atp_proof =
blanchet@50898
   951
          atp_proof
blanchet@50898
   952
          |> clean_up_atp_proof_dependencies
blanchet@50898
   953
          |> nasty_atp_proof pool
blanchet@50898
   954
          |> map_term_names_in_atp_proof repair_name
blanchet@50898
   955
          |> decode_lines ctxt sym_tab
blanchet@50898
   956
          |> rpair [] |-> fold_rev (add_line fact_names)
blanchet@50898
   957
          |> repair_waldmeister_endgame
blanchet@50898
   958
          |> rpair [] |-> fold_rev add_nontrivial_line
blanchet@50898
   959
          |> rpair (0, [])
blanchet@50933
   960
          |-> fold_rev (add_desired_line fact_names frees)
blanchet@50898
   961
          |> snd
blanchet@50898
   962
        val conj_name = conjecture_prefix ^ string_of_int (length hyp_ts)
blanchet@50898
   963
        val conjs =
blanchet@51025
   964
          atp_proof |> map_filter
blanchet@51027
   965
            (fn Inference_Step (name as (_, ss), _, _, _, []) =>
blanchet@51025
   966
                if member (op =) ss conj_name then SOME name else NONE
blanchet@51025
   967
              | _ => NONE)
blanchet@51025
   968
        val assms =
blanchet@51025
   969
          atp_proof |> map_filter
blanchet@51028
   970
            (fn Inference_Step (name as (_, ss), _, _, _, []) =>
blanchet@51028
   971
                (case resolve_conjecture ss of
blanchet@51028
   972
                   [j] =>
blanchet@51028
   973
                   if j = length hyp_ts then NONE
blanchet@51028
   974
                   else SOME (Assume (raw_label_for_name name, nth hyp_ts j))
blanchet@51028
   975
                 | _ => NONE)
blanchet@51025
   976
              | _ => NONE)
blanchet@50898
   977
        fun dep_of_step (Definition_Step _) = NONE
blanchet@51027
   978
          | dep_of_step (Inference_Step (name, _, _, _, from)) =
blanchet@51027
   979
            SOME (from, name)
blanchet@50898
   980
        val ref_graph = atp_proof |> map_filter dep_of_step |> make_ref_graph
blanchet@50898
   981
        val axioms = axioms_of_ref_graph ref_graph conjs
blanchet@50898
   982
        val tainted = tainted_atoms_of_ref_graph ref_graph conjs
blanchet@50898
   983
        val props =
blanchet@50898
   984
          Symtab.empty
blanchet@50898
   985
          |> fold (fn Definition_Step _ => I (* FIXME *)
blanchet@51028
   986
                    | Inference_Step (name as (s, ss), role, t, _, _) =>
blanchet@50898
   987
                      Symtab.update_new (s,
blanchet@51031
   988
                        if member (op = o apsnd fst) tainted s then
blanchet@51031
   989
                          t |> role <> Conjecture ? s_not
blanchet@51031
   990
                            |> fold exists_of (map Var (Term.add_vars t []))
blanchet@51031
   991
                        else
blanchet@51031
   992
                          t))
blanchet@50898
   993
                  atp_proof
blanchet@51028
   994
        (* The assumptions and conjecture are props; the rest are bools. *)
blanchet@51031
   995
        fun prop_of_clause [name as (s, ss)] =
blanchet@51031
   996
            (case resolve_conjecture ss of
blanchet@51031
   997
               [j] => if j = length hyp_ts then concl_t else nth hyp_ts j
blanchet@51031
   998
             | _ => the_default @{term False} (Symtab.lookup props s)
blanchet@51031
   999
                    |> HOLogic.mk_Trueprop |> close_form)
blanchet@51031
  1000
          | prop_of_clause names =
blanchet@51033
  1001
            let val lits = map_filter (Symtab.lookup props o fst) names in
blanchet@51033
  1002
              case List.partition (can HOLogic.dest_not) lits of
blanchet@51033
  1003
                (negs as _ :: _, pos as _ :: _) =>
blanchet@51033
  1004
                HOLogic.mk_imp
blanchet@51033
  1005
                  (Library.foldr1 s_conj (map HOLogic.dest_not negs),
blanchet@51033
  1006
                   Library.foldr1 s_disj pos)
blanchet@51033
  1007
              | _ => fold (curry s_disj) lits @{term False}
blanchet@51033
  1008
            end
blanchet@51031
  1009
            |> HOLogic.mk_Trueprop |> close_form
blanchet@50898
  1010
        fun maybe_show outer c =
blanchet@50898
  1011
          (outer andalso length c = 1 andalso subset (op =) (c, conjs))
blanchet@50898
  1012
          ? cons Show
blanchet@50898
  1013
        fun do_have outer qs (gamma, c) =
blanchet@50898
  1014
          Prove (maybe_show outer c qs, label_of_clause c, prop_of_clause c,
blanchet@51020
  1015
                 By_Metis (fold (add_fact_from_dependencies fact_names) gamma
blanchet@51020
  1016
                                ([], [])))
blanchet@50898
  1017
        fun do_inf outer (Have z) = do_have outer [] z
blanchet@50898
  1018
          | do_inf outer (Cases cases) =
blanchet@50898
  1019
            let val c = succedent_of_cases cases in
blanchet@50898
  1020
              Prove (maybe_show outer c [Ultimately], label_of_clause c,
blanchet@50898
  1021
                     prop_of_clause c,
blanchet@50898
  1022
                     Case_Split (map (do_case false) cases, ([], [])))
blanchet@50898
  1023
            end
blanchet@50898
  1024
        and do_case outer (c, infs) =
blanchet@50898
  1025
          Assume (label_of_clause c, prop_of_clause c) ::
blanchet@50898
  1026
          map (do_inf outer) infs
blanchet@51019
  1027
        val (ext_time, isar_proof) =
blanchet@51019
  1028
          ref_graph
blanchet@51019
  1029
          |> redirect_graph axioms tainted
blanchet@51019
  1030
          |> map (do_inf true)
blanchet@51025
  1031
          |> append assms
blanchet@51019
  1032
          |> (if isar_shrinkage <= 1.0 andalso isar_proofs then
blanchet@51019
  1033
                pair (true, seconds 0.0)
blanchet@51019
  1034
              else
blanchet@51019
  1035
                shrink_proof debug ctxt type_enc lam_trans preplay
blanchet@51019
  1036
                     preplay_timeout
blanchet@51019
  1037
                     (if isar_proofs then isar_shrinkage else 1000.0))
blanchet@51019
  1038
       (* ||> reorder_proof_to_minimize_jumps (* ? *) *)
blanchet@51019
  1039
          ||> chain_direct_proof
blanchet@51019
  1040
          ||> kill_useless_labels_in_proof
blanchet@51019
  1041
          ||> relabel_proof
blanchet@51019
  1042
          ||> not (null params) ? cons (Fix params)
blanchet@50933
  1043
        val num_steps = length isar_proof
blanchet@50933
  1044
        val isar_text =
blanchet@50933
  1045
          string_for_proof ctxt type_enc lam_trans subgoal subgoal_count
blanchet@50933
  1046
                           isar_proof
blanchet@50898
  1047
      in
blanchet@50933
  1048
        case isar_text of
blanchet@50898
  1049
          "" =>
blanchet@50933
  1050
          if isar_proofs then
blanchet@50898
  1051
            "\nNo structured proof available (proof too short)."
blanchet@50898
  1052
          else
blanchet@50898
  1053
            ""
blanchet@50898
  1054
        | _ =>
blanchet@50933
  1055
          "\n\n" ^
blanchet@50933
  1056
          (if isar_proofs then
blanchet@50933
  1057
             "Structured proof" ^
blanchet@50933
  1058
             (if verbose then
blanchet@50933
  1059
                " (" ^ string_of_int num_steps ^ " step" ^ plural_s num_steps ^
blanchet@51019
  1060
                (if preplay andalso isar_shrinkage > 1.0 then
blanchet@51019
  1061
                   ", " ^ string_from_ext_time ext_time
blanchet@51019
  1062
                 else
blanchet@51019
  1063
                   "") ^
blanchet@50933
  1064
                ")"
blanchet@50933
  1065
              else
blanchet@50933
  1066
                "")
blanchet@50933
  1067
           else
blanchet@50933
  1068
             "Perhaps this will work") ^
blanchet@50933
  1069
          ":\n" ^ Markup.markup Isabelle_Markup.sendback isar_text
blanchet@50898
  1070
      end
blanchet@50898
  1071
    val isar_proof =
blanchet@50898
  1072
      if debug then
blanchet@50898
  1073
        isar_proof_of ()
blanchet@50898
  1074
      else case try isar_proof_of () of
blanchet@50898
  1075
        SOME s => s
blanchet@50933
  1076
      | NONE => if isar_proofs then
blanchet@50898
  1077
                  "\nWarning: The Isar proof construction failed."
blanchet@50898
  1078
                else
blanchet@50898
  1079
                  ""
blanchet@50898
  1080
  in one_line_proof ^ isar_proof end
blanchet@50898
  1081
blanchet@50933
  1082
fun proof_text ctxt isar_proofs isar_params num_chained
blanchet@50898
  1083
               (one_line_params as (preplay, _, _, _, _, _)) =
blanchet@50933
  1084
  (if case preplay of Failed_to_Play _ => true | _ => isar_proofs then
blanchet@50933
  1085
     isar_proof_text ctxt isar_proofs isar_params
blanchet@50898
  1086
   else
blanchet@50898
  1087
     one_line_proof_text num_chained) one_line_params
blanchet@50898
  1088
blanchet@50898
  1089
end;