src/HOL/Tools/ATP/atp_reconstruct.ML
author blanchet
Mon, 06 Jun 2011 20:36:35 +0200
changeset 44023 649bada59658
parent 44017 29a3a1a7794d
child 44024 faece9668bce
permissions -rw-r--r--
slacker version of "fastype_of", in case a function has dummy type
blanchet@43926
     1
(*  Title:      HOL/Tools/ATP/atp_reconstruct.ML
blanchet@38261
     2
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
blanchet@38261
     3
    Author:     Claire Quigley, Cambridge University Computer Laboratory
blanchet@36392
     4
    Author:     Jasmin Blanchette, TU Muenchen
paulson@21978
     5
blanchet@43926
     6
Proof reconstruction from ATP proofs.
wenzelm@33318
     7
*)
wenzelm@33318
     8
blanchet@43926
     9
signature ATP_RECONSTRUCT =
paulson@24425
    10
sig
blanchet@43935
    11
  type 'a fo_term = 'a ATP_Problem.fo_term
blanchet@43968
    12
  type ('a, 'b, 'c) formula = ('a, 'b, 'c) ATP_Problem.formula
blanchet@43320
    13
  type 'a proof = 'a ATP_Proof.proof
blanchet@43926
    14
  type locality = ATP_Translate.locality
blanchet@43943
    15
  type type_sys = ATP_Translate.type_sys
blanchet@43874
    16
blanchet@43874
    17
  datatype reconstructor =
blanchet@43874
    18
    Metis |
blanchet@43874
    19
    MetisFT |
blanchet@44009
    20
    MetisX |
blanchet@43874
    21
    SMT of string
blanchet@43874
    22
blanchet@43891
    23
  datatype play =
blanchet@43891
    24
    Played of reconstructor * Time.time |
blanchet@43874
    25
    Trust_Playable of reconstructor * Time.time option|
blanchet@44007
    26
    Failed_to_Play of reconstructor
blanchet@43874
    27
blanchet@36281
    28
  type minimize_command = string list -> string
blanchet@43874
    29
  type one_line_params =
blanchet@43891
    30
    play * string * (string * locality) list * minimize_command * int * int
blanchet@39053
    31
  type isar_params =
blanchet@43943
    32
    bool * bool * int * type_sys * string Symtab.table * int list list * int
blanchet@43943
    33
    * (string * locality) list vector * int Symtab.table * string proof * thm
blanchet@40445
    34
  val repair_conjecture_shape_and_fact_names :
blanchet@43943
    35
    type_sys -> string -> int list list -> int
blanchet@43750
    36
    -> (string * locality) list vector -> int list
blanchet@43750
    37
    -> int list list * int * (string * locality) list vector * int list
blanchet@43322
    38
  val used_facts_in_atp_proof :
blanchet@43943
    39
    Proof.context -> type_sys -> int -> (string * locality) list vector
blanchet@43809
    40
    -> string proof -> (string * locality) list
blanchet@43745
    41
  val used_facts_in_unsound_atp_proof :
blanchet@43943
    42
    Proof.context -> type_sys -> int list list -> int
blanchet@43874
    43
    -> (string * locality) list vector -> 'a proof -> string list option
blanchet@43874
    44
  val uses_typed_helpers : int list -> 'a proof -> bool
blanchet@43892
    45
  val reconstructor_name : reconstructor -> string
blanchet@43874
    46
  val one_line_proof_text : one_line_params -> string
blanchet@43976
    47
  val make_tvar : string -> typ
blanchet@43976
    48
  val make_tfree : Proof.context -> string -> typ
blanchet@43935
    49
  val term_from_atp :
blanchet@43976
    50
    Proof.context -> bool -> int Symtab.table -> typ option -> string fo_term
blanchet@43976
    51
    -> term
blanchet@43968
    52
  val prop_from_atp :
blanchet@43976
    53
    Proof.context -> bool -> int Symtab.table
blanchet@43968
    54
    -> (string, string, string fo_term) formula -> term
blanchet@43809
    55
  val isar_proof_text :
blanchet@43903
    56
    Proof.context -> bool -> isar_params -> one_line_params -> string
blanchet@43809
    57
  val proof_text :
blanchet@43874
    58
    Proof.context -> bool -> isar_params -> one_line_params -> string
paulson@24425
    59
end;
paulson@21978
    60
blanchet@43926
    61
structure ATP_Reconstruct : ATP_RECONSTRUCT =
paulson@21978
    62
struct
paulson@21978
    63
blanchet@43926
    64
open ATP_Util
blanchet@38262
    65
open ATP_Problem
blanchet@39692
    66
open ATP_Proof
blanchet@43926
    67
open ATP_Translate
blanchet@35826
    68
blanchet@43874
    69
datatype reconstructor =
blanchet@43874
    70
  Metis |
blanchet@43874
    71
  MetisFT |
blanchet@44009
    72
  MetisX |
blanchet@43874
    73
  SMT of string
blanchet@43874
    74
blanchet@43891
    75
datatype play =
blanchet@43891
    76
  Played of reconstructor * Time.time |
blanchet@43874
    77
  Trust_Playable of reconstructor * Time.time option |
blanchet@44007
    78
  Failed_to_Play of reconstructor
blanchet@43874
    79
blanchet@36281
    80
type minimize_command = string list -> string
blanchet@43874
    81
type one_line_params =
blanchet@43891
    82
  play * string * (string * locality) list * minimize_command * int * int
blanchet@39053
    83
type isar_params =
blanchet@43943
    84
  bool * bool * int * type_sys * string Symtab.table * int list list * int
blanchet@43878
    85
  * (string * locality) list vector * int Symtab.table * string proof * thm
blanchet@36281
    86
blanchet@39740
    87
fun is_head_digit s = Char.isDigit (String.sub (s, 0))
blanchet@39740
    88
val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode)
blanchet@39740
    89
blanchet@43750
    90
val is_typed_helper_name =
blanchet@43750
    91
  String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
blanchet@43750
    92
blanchet@39740
    93
fun find_first_in_list_vector vec key =
blanchet@39740
    94
  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
blanchet@39740
    95
                 | (_, value) => value) NONE vec
blanchet@39740
    96
blanchet@39693
    97
blanchet@43880
    98
(** SPASS's FLOTTER hack **)
blanchet@39733
    99
blanchet@40445
   100
(* This is a hack required for keeping track of facts after they have been
blanchet@43880
   101
   clausified by SPASS's FLOTTER preprocessor. The "ATP/scripts/spass" script is
blanchet@43880
   102
   also part of this hack. *)
blanchet@39733
   103
blanchet@39733
   104
val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
blanchet@39733
   105
blanchet@39733
   106
fun extract_clause_sequence output =
blanchet@39733
   107
  let
blanchet@39733
   108
    val tokens_of = String.tokens (not o Char.isAlphaNum)
blanchet@43458
   109
    fun extract_num ("clause" :: (ss as _ :: _)) = Int.fromString (List.last ss)
blanchet@39733
   110
      | extract_num _ = NONE
blanchet@39733
   111
  in output |> split_lines |> map_filter (extract_num o tokens_of) end
blanchet@39733
   112
blanchet@39733
   113
val parse_clause_formula_pair =
blanchet@39733
   114
  $$ "(" |-- scan_integer --| $$ ","
blanchet@39733
   115
  -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
blanchet@39733
   116
  --| Scan.option ($$ ",")
blanchet@39733
   117
val parse_clause_formula_relation =
blanchet@39733
   118
  Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
blanchet@39733
   119
  |-- Scan.repeat parse_clause_formula_pair
blanchet@39733
   120
val extract_clause_formula_relation =
blanchet@39733
   121
  Substring.full #> Substring.position set_ClauseFormulaRelationN
blanchet@39733
   122
  #> snd #> Substring.position "." #> fst #> Substring.string
wenzelm@40875
   123
  #> raw_explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
blanchet@39733
   124
  #> fst
blanchet@39733
   125
blanchet@43518
   126
fun maybe_unprefix_fact_number type_sys =
blanchet@43518
   127
  polymorphism_of_type_sys type_sys <> Polymorphic
blanchet@43518
   128
  ? (space_implode "_" o tl o space_explode "_")
blanchet@43051
   129
blanchet@43518
   130
fun repair_conjecture_shape_and_fact_names type_sys output conjecture_shape
blanchet@43750
   131
        fact_offset fact_names typed_helpers =
blanchet@39733
   132
  if String.isSubstring set_ClauseFormulaRelationN output then
blanchet@39733
   133
    let
blanchet@39733
   134
      val j0 = hd (hd conjecture_shape)
blanchet@39733
   135
      val seq = extract_clause_sequence output
blanchet@39733
   136
      val name_map = extract_clause_formula_relation output
blanchet@39733
   137
      fun renumber_conjecture j =
blanchet@39733
   138
        conjecture_prefix ^ string_of_int (j - j0)
blanchet@39733
   139
        |> AList.find (fn (s, ss) => member (op =) ss s) name_map
blanchet@39733
   140
        |> map (fn s => find_index (curry (op =) s) seq + 1)
blanchet@39733
   141
      fun names_for_number j =
blanchet@39733
   142
        j |> AList.lookup (op =) name_map |> these
blanchet@43518
   143
          |> map_filter (try (unascii_of o maybe_unprefix_fact_number type_sys
blanchet@43051
   144
                              o unprefix fact_prefix))
blanchet@39733
   145
          |> map (fn name =>
blanchet@40445
   146
                     (name, name |> find_first_in_list_vector fact_names |> the)
blanchet@39733
   147
                     handle Option.Option =>
blanchet@39733
   148
                            error ("No such fact: " ^ quote name ^ "."))
blanchet@39733
   149
    in
blanchet@43458
   150
      (conjecture_shape |> map (maps renumber_conjecture), 0,
blanchet@43750
   151
       seq |> map names_for_number |> Vector.fromList,
blanchet@43750
   152
       name_map |> filter (forall is_typed_helper_name o snd) |> map fst)
blanchet@39733
   153
    end
blanchet@39733
   154
  else
blanchet@43750
   155
    (conjecture_shape, fact_offset, fact_names, typed_helpers)
blanchet@39733
   156
blanchet@43051
   157
val vampire_step_prefix = "f" (* grrr... *)
blanchet@41451
   158
blanchet@43750
   159
val extract_step_number =
blanchet@43750
   160
  Int.fromString o perhaps (try (unprefix vampire_step_prefix))
blanchet@43750
   161
blanchet@43616
   162
fun resolve_fact type_sys _ fact_names (_, SOME s) =
blanchet@43051
   163
    (case try (unprefix fact_prefix) s of
blanchet@43051
   164
       SOME s' =>
blanchet@43518
   165
       let val s' = s' |> maybe_unprefix_fact_number type_sys |> unascii_of in
blanchet@43051
   166
         case find_first_in_list_vector fact_names s' of
blanchet@43051
   167
           SOME x => [(s', x)]
blanchet@43051
   168
         | NONE => []
blanchet@43051
   169
       end
blanchet@39693
   170
     | NONE => [])
blanchet@43518
   171
  | resolve_fact _ facts_offset fact_names (num, NONE) =
blanchet@43750
   172
    (case extract_step_number num of
blanchet@43750
   173
       SOME j =>
blanchet@43750
   174
       let val j = j - facts_offset in
blanchet@43750
   175
         if j > 0 andalso j <= Vector.length fact_names then
blanchet@43750
   176
           Vector.sub (fact_names, j - 1)
blanchet@43750
   177
         else
blanchet@43750
   178
           []
blanchet@43750
   179
       end
blanchet@43750
   180
     | NONE => [])
blanchet@43750
   181
blanchet@43750
   182
fun is_fact type_sys conjecture_shape =
blanchet@43750
   183
  not o null o resolve_fact type_sys 0 conjecture_shape
blanchet@39693
   184
blanchet@43616
   185
fun resolve_conjecture _ (_, SOME s) =
blanchet@43616
   186
    (case try (unprefix conjecture_prefix) s of
blanchet@43616
   187
       SOME s' =>
blanchet@43616
   188
       (case Int.fromString s' of
blanchet@43616
   189
          SOME j => [j]
blanchet@43616
   190
        | NONE => [])
blanchet@43616
   191
     | NONE => [])
blanchet@43616
   192
  | resolve_conjecture conjecture_shape (num, NONE) =
blanchet@43750
   193
    case extract_step_number num of
blanchet@43750
   194
      SOME i => (case find_index (exists (curry (op =) i)) conjecture_shape of
blanchet@43750
   195
                   ~1 => []
blanchet@43750
   196
                 | j => [j])
blanchet@43616
   197
    | NONE => []
blanchet@43320
   198
blanchet@43320
   199
fun is_conjecture conjecture_shape =
blanchet@43320
   200
  not o null o resolve_conjecture conjecture_shape
blanchet@43320
   201
blanchet@43750
   202
fun is_typed_helper _ (_, SOME s) = is_typed_helper_name s
blanchet@43750
   203
  | is_typed_helper typed_helpers (num, NONE) =
blanchet@43750
   204
    (case extract_step_number num of
blanchet@43750
   205
       SOME i => member (op =) typed_helpers i
blanchet@43750
   206
     | NONE => false)
blanchet@43750
   207
blanchet@43809
   208
val leo2_ext = "extcnf_equal_neg"
blanchet@43809
   209
val isa_ext = Thm.get_name_hint @{thm ext}
blanchet@43809
   210
val isa_short_ext = Long_Name.base_name isa_ext
blanchet@43809
   211
blanchet@43809
   212
fun ext_name ctxt =
blanchet@43809
   213
  if Thm.eq_thm_prop (@{thm ext},
blanchet@43809
   214
         singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
blanchet@43809
   215
    isa_short_ext
blanchet@43809
   216
  else
blanchet@43809
   217
    isa_ext
blanchet@43809
   218
blanchet@43809
   219
fun add_fact _ type_sys facts_offset fact_names (Inference (name, _, [])) =
blanchet@43784
   220
    union (op =) (resolve_fact type_sys facts_offset fact_names name)
blanchet@43809
   221
  | add_fact ctxt _ _ _ (Inference (_, _, deps)) =
blanchet@43809
   222
    if AList.defined (op =) deps leo2_ext then
blanchet@43809
   223
      insert (op =) (ext_name ctxt, General (* or Chained... *))
blanchet@43809
   224
    else
blanchet@43809
   225
      I
blanchet@43809
   226
  | add_fact _ _ _ _ _ = I
blanchet@39693
   227
blanchet@43809
   228
fun used_facts_in_atp_proof ctxt type_sys facts_offset fact_names atp_proof =
blanchet@43320
   229
  if null atp_proof then Vector.foldl (op @) [] fact_names
blanchet@43809
   230
  else fold (add_fact ctxt type_sys facts_offset fact_names) atp_proof []
blanchet@43320
   231
blanchet@43320
   232
fun is_conjecture_referred_to_in_proof conjecture_shape =
blanchet@43320
   233
  exists (fn Inference (name, _, []) => is_conjecture conjecture_shape name
blanchet@43320
   234
           | _ => false)
blanchet@43320
   235
blanchet@43809
   236
fun used_facts_in_unsound_atp_proof ctxt type_sys conjecture_shape facts_offset
blanchet@43745
   237
                                    fact_names atp_proof =
blanchet@43745
   238
  let
blanchet@43809
   239
    val used_facts =
blanchet@43809
   240
      used_facts_in_atp_proof ctxt type_sys facts_offset fact_names atp_proof
blanchet@43745
   241
  in
blanchet@43745
   242
    if forall (is_locality_global o snd) used_facts andalso
blanchet@43745
   243
       not (is_conjecture_referred_to_in_proof conjecture_shape atp_proof) then
blanchet@43745
   244
      SOME (map fst used_facts)
blanchet@43745
   245
    else
blanchet@43745
   246
      NONE
blanchet@43745
   247
  end
blanchet@43320
   248
blanchet@43874
   249
fun uses_typed_helpers typed_helpers =
blanchet@43874
   250
  exists (fn Inference (name, _, []) => is_typed_helper typed_helpers name
blanchet@43874
   251
           | _ => false)
blanchet@43874
   252
blanchet@43874
   253
blanchet@43320
   254
(** Soft-core proof reconstruction: Metis one-liner **)
blanchet@43320
   255
blanchet@43874
   256
fun reconstructor_name Metis = "metis"
blanchet@43874
   257
  | reconstructor_name MetisFT = "metisFT"
blanchet@44009
   258
  | reconstructor_name MetisX = "metisX"
blanchet@43874
   259
  | reconstructor_name (SMT _) = "smt"
blanchet@43874
   260
blanchet@43874
   261
fun reconstructor_settings (SMT settings) = settings
blanchet@43874
   262
  | reconstructor_settings _ = ""
blanchet@43874
   263
blanchet@43320
   264
fun string_for_label (s, num) = s ^ string_of_int num
blanchet@43320
   265
blanchet@43874
   266
fun show_time NONE = ""
blanchet@43875
   267
  | show_time (SOME ext_time) = " (" ^ string_from_ext_time ext_time ^ ")"
blanchet@43874
   268
blanchet@43320
   269
fun set_settings "" = ""
blanchet@43320
   270
  | set_settings settings = "using [[" ^ settings ^ "]] "
blanchet@43320
   271
fun apply_on_subgoal settings _ 1 = set_settings settings ^ "by "
blanchet@43320
   272
  | apply_on_subgoal settings 1 _ = set_settings settings ^ "apply "
blanchet@43320
   273
  | apply_on_subgoal settings i n =
blanchet@43320
   274
    "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal settings 1 n
blanchet@43320
   275
fun command_call name [] = name
blanchet@43320
   276
  | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
blanchet@43874
   277
fun try_command_line banner time command =
blanchet@43874
   278
  banner ^ ": " ^ Markup.markup Markup.sendback command ^ show_time time ^ "."
blanchet@43320
   279
fun using_labels [] = ""
blanchet@43320
   280
  | using_labels ls =
blanchet@43320
   281
    "using " ^ space_implode " " (map string_for_label ls) ^ " "
blanchet@43874
   282
fun reconstructor_command reconstructor i n (ls, ss) =
blanchet@43874
   283
  using_labels ls ^
blanchet@43874
   284
  apply_on_subgoal (reconstructor_settings reconstructor) i n ^
blanchet@43874
   285
  command_call (reconstructor_name reconstructor) ss
blanchet@43320
   286
fun minimize_line _ [] = ""
blanchet@43320
   287
  | minimize_line minimize_command ss =
blanchet@43320
   288
    case minimize_command ss of
blanchet@43320
   289
      "" => ""
blanchet@43847
   290
    | command => "\nTo minimize: " ^ Markup.markup Markup.sendback command ^ "."
blanchet@39693
   291
blanchet@40966
   292
val split_used_facts =
blanchet@40966
   293
  List.partition (curry (op =) Chained o snd)
blanchet@39693
   294
  #> pairself (sort_distinct (string_ord o pairself fst))
blanchet@39693
   295
blanchet@43874
   296
fun one_line_proof_text (preplay, banner, used_facts, minimize_command,
blanchet@43878
   297
                         subgoal, subgoal_count) =
blanchet@39693
   298
  let
blanchet@43874
   299
    val (chained, extra) = split_used_facts used_facts
blanchet@44007
   300
    val (failed, reconstructor, ext_time) =
blanchet@43874
   301
      case preplay of
blanchet@43891
   302
        Played (reconstructor, time) =>
blanchet@44007
   303
        (false, reconstructor, (SOME (false, time)))
blanchet@43874
   304
      | Trust_Playable (reconstructor, time) =>
blanchet@44007
   305
        (false, reconstructor,
blanchet@43874
   306
         case time of
blanchet@43874
   307
           NONE => NONE
blanchet@43874
   308
         | SOME time =>
blanchet@43874
   309
           if time = Time.zeroTime then NONE else SOME (true, time))
blanchet@44007
   310
      | Failed_to_Play reconstructor => (true, reconstructor, NONE)
blanchet@43874
   311
    val try_line =
blanchet@44007
   312
      ([], map fst extra)
blanchet@44007
   313
      |> reconstructor_command reconstructor subgoal subgoal_count
blanchet@44007
   314
      |> (if failed then enclose "One-line proof reconstruction failed: " "."
blanchet@44007
   315
          else try_command_line banner ext_time)
blanchet@43874
   316
  in try_line ^ minimize_line minimize_command (map fst (extra @ chained)) end
blanchet@39693
   317
blanchet@39693
   318
(** Hard-core proof reconstruction: structured Isar proofs **)
blanchet@39693
   319
blanchet@38248
   320
(* Simple simplifications to ensure that sort annotations don't leave a trail of
blanchet@38248
   321
   spurious "True"s. *)
blanchet@43477
   322
fun s_not (Const (@{const_name All}, T) $ Abs (s, T', t')) =
blanchet@43477
   323
    Const (@{const_name Ex}, T) $ Abs (s, T', s_not t')
blanchet@43477
   324
  | s_not (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
blanchet@43477
   325
    Const (@{const_name All}, T) $ Abs (s, T', s_not t')
blanchet@43477
   326
  | s_not (@{const HOL.implies} $ t1 $ t2) = @{const HOL.conj} $ t1 $ s_not t2
blanchet@43477
   327
  | s_not (@{const HOL.conj} $ t1 $ t2) =
blanchet@43477
   328
    @{const HOL.disj} $ s_not t1 $ s_not t2
blanchet@43477
   329
  | s_not (@{const HOL.disj} $ t1 $ t2) =
blanchet@43477
   330
    @{const HOL.conj} $ s_not t1 $ s_not t2
blanchet@43477
   331
  | s_not (@{const False}) = @{const True}
blanchet@43477
   332
  | s_not (@{const True}) = @{const False}
blanchet@38248
   333
  | s_not (@{const Not} $ t) = t
blanchet@38248
   334
  | s_not t = @{const Not} $ t
blanchet@38248
   335
fun s_conj (@{const True}, t2) = t2
blanchet@38248
   336
  | s_conj (t1, @{const True}) = t1
blanchet@38248
   337
  | s_conj p = HOLogic.mk_conj p
blanchet@38248
   338
fun s_disj (@{const False}, t2) = t2
blanchet@38248
   339
  | s_disj (t1, @{const False}) = t1
blanchet@38248
   340
  | s_disj p = HOLogic.mk_disj p
blanchet@38248
   341
fun s_imp (@{const True}, t2) = t2
blanchet@38248
   342
  | s_imp (t1, @{const False}) = s_not t1
blanchet@38248
   343
  | s_imp p = HOLogic.mk_imp p
blanchet@38248
   344
fun s_iff (@{const True}, t2) = t2
blanchet@38248
   345
  | s_iff (t1, @{const True}) = t1
blanchet@38248
   346
  | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
blanchet@38248
   347
blanchet@39671
   348
fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
blanchet@39671
   349
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t
blanchet@39671
   350
blanchet@43976
   351
fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
blanchet@43976
   352
fun make_tfree ctxt w =
blanchet@43976
   353
  let val ww = "'" ^ w in
blanchet@43976
   354
    TFree (ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))
blanchet@43976
   355
  end
blanchet@43976
   356
blanchet@39693
   357
val indent_size = 2
blanchet@39693
   358
val no_label = ("", ~1)
blanchet@39693
   359
blanchet@39693
   360
val raw_prefix = "X"
blanchet@39693
   361
val assum_prefix = "A"
blanchet@43051
   362
val have_prefix = "F"
blanchet@39616
   363
blanchet@39693
   364
fun raw_label_for_name conjecture_shape name =
blanchet@39693
   365
  case resolve_conjecture conjecture_shape name of
blanchet@39693
   366
    [j] => (conjecture_prefix, j)
blanchet@39695
   367
  | _ => case Int.fromString (fst name) of
blanchet@39693
   368
           SOME j => (raw_prefix, j)
blanchet@39695
   369
         | NONE => (raw_prefix ^ fst name, 0)
blanchet@39693
   370
paulson@21978
   371
(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
paulson@21978
   372
blanchet@38225
   373
exception FO_TERM of string fo_term list
blanchet@43402
   374
exception FORMULA of (string, string, string fo_term) formula list
blanchet@38225
   375
exception SAME of unit
paulson@21978
   376
blanchet@36901
   377
(* Type variables are given the basic sort "HOL.type". Some will later be
blanchet@38225
   378
   constrained by information from type literals, or by type inference. *)
blanchet@43976
   379
fun typ_from_atp ctxt (u as ATerm (a, us)) =
blanchet@43976
   380
  let val Ts = map (typ_from_atp ctxt) us in
blanchet@38987
   381
    case strip_prefix_and_unascii type_const_prefix a of
blanchet@38225
   382
      SOME b => Type (invert_const b, Ts)
blanchet@38225
   383
    | NONE =>
blanchet@38225
   384
      if not (null us) then
blanchet@38225
   385
        raise FO_TERM [u]  (* only "tconst"s have type arguments *)
blanchet@38987
   386
      else case strip_prefix_and_unascii tfree_prefix a of
blanchet@43976
   387
        SOME b => make_tfree ctxt b
blanchet@36482
   388
      | NONE =>
blanchet@38987
   389
        case strip_prefix_and_unascii tvar_prefix a of
blanchet@43976
   390
          SOME b => make_tvar b
blanchet@36482
   391
        | NONE =>
blanchet@38225
   392
          (* Variable from the ATP, say "X1" *)
blanchet@38225
   393
          Type_Infer.param 0 (a, HOLogic.typeS)
blanchet@38225
   394
  end
paulson@21978
   395
blanchet@38248
   396
(* Type class literal applied to a type. Returns triple of polarity, class,
blanchet@38248
   397
   type. *)
blanchet@43976
   398
fun type_constraint_from_term ctxt (u as ATerm (a, us)) =
blanchet@43976
   399
  case (strip_prefix_and_unascii class_prefix a, map (typ_from_atp ctxt) us) of
blanchet@43477
   400
    (SOME b, [T]) => (b, T)
blanchet@38248
   401
  | _ => raise FO_TERM [u]
blanchet@38248
   402
blanchet@38331
   403
(** Accumulate type constraints in a formula: negative type literals **)
blanchet@38248
   404
fun add_var (key, z)  = Vartab.map_default (key, []) (cons z)
blanchet@43477
   405
fun add_type_constraint false (cl, TFree (a ,_)) = add_var ((a, ~1), cl)
blanchet@43477
   406
  | add_type_constraint false (cl, TVar (ix, _)) = add_var (ix, cl)
blanchet@43477
   407
  | add_type_constraint _ _ = I
blanchet@38248
   408
blanchet@43935
   409
fun repair_variable_name f s =
blanchet@36482
   410
  let
blanchet@36482
   411
    fun subscript_name s n = s ^ nat_subscript n
blanchet@38713
   412
    val s = String.map f s
blanchet@36482
   413
  in
blanchet@36482
   414
    case space_explode "_" s of
blanchet@36482
   415
      [_] => (case take_suffix Char.isDigit (String.explode s) of
blanchet@36482
   416
                (cs1 as _ :: _, cs2 as _ :: _) =>
blanchet@36482
   417
                subscript_name (String.implode cs1)
blanchet@36482
   418
                               (the (Int.fromString (String.implode cs2)))
blanchet@36482
   419
              | (_, _) => s)
blanchet@36482
   420
    | [s1, s2] => (case Int.fromString s2 of
blanchet@36482
   421
                     SOME n => subscript_name s1 n
blanchet@36482
   422
                   | NONE => s)
blanchet@36482
   423
    | _ => s
blanchet@36482
   424
  end
blanchet@44023
   425
blanchet@44023
   426
fun slack_fastype_of t = fastype_of t handle TERM _ => HOLogic.typeT
blanchet@44023
   427
blanchet@36901
   428
(* First-order translation. No types are known for variables. "HOLogic.typeT"
blanchet@38248
   429
   should allow them to be inferred. *)
blanchet@43976
   430
fun term_from_atp ctxt textual sym_tab =
blanchet@36901
   431
  let
blanchet@43976
   432
    val thy = Proof_Context.theory_of ctxt
blanchet@43935
   433
    (* see also "mk_var" in "Metis_Reconstruct" *)
blanchet@43935
   434
    val var_index = if textual then 0 else 1
blanchet@43972
   435
    fun do_term extra_ts opt_T u =
blanchet@36901
   436
      case u of
blanchet@43410
   437
        ATerm (a, us) =>
blanchet@43803
   438
        if String.isPrefix simple_type_prefix a then
blanchet@43803
   439
          @{const True} (* ignore TPTP type information *)
blanchet@43841
   440
        else if a = tptp_equal then
blanchet@43934
   441
          let val ts = map (do_term [] NONE) us in
blanchet@43935
   442
            if textual andalso length ts = 2 andalso
blanchet@43935
   443
              hd ts aconv List.last ts then
blanchet@39352
   444
              (* Vampire is keen on producing these. *)
blanchet@39352
   445
              @{const True}
blanchet@39352
   446
            else
blanchet@39352
   447
              list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
blanchet@39352
   448
          end
blanchet@43841
   449
        else case strip_prefix_and_unascii const_prefix a of
blanchet@43841
   450
          SOME s =>
blanchet@43626
   451
          let
blanchet@43626
   452
            val ((s', s), mangled_us) = s |> unmangled_const |>> `invert_const
blanchet@43626
   453
          in
blanchet@43620
   454
            if s' = type_tag_name then
blanchet@43460
   455
              case mangled_us @ us of
blanchet@43460
   456
                [typ_u, term_u] =>
blanchet@43976
   457
                do_term extra_ts (SOME (typ_from_atp ctxt typ_u)) term_u
blanchet@43460
   458
              | _ => raise FO_TERM us
blanchet@43807
   459
            else if s' = predicator_name then
blanchet@43934
   460
              do_term [] (SOME @{typ bool}) (hd us)
blanchet@43807
   461
            else if s' = app_op_name then
blanchet@43972
   462
              let val extra_t = do_term [] NONE (List.last us) in
blanchet@43972
   463
                do_term (extra_t :: extra_ts)
blanchet@43972
   464
                        (case opt_T of
blanchet@44023
   465
                           SOME T => SOME (slack_fastype_of extra_t --> T)
blanchet@43972
   466
                         | NONE => NONE)
blanchet@43972
   467
                        (nth us (length us - 2))
blanchet@43972
   468
              end
blanchet@43807
   469
            else if s' = type_pred_name then
blanchet@43422
   470
              @{const True} (* ignore type predicates *)
blanchet@43420
   471
            else
blanchet@43420
   472
              let
blanchet@43620
   473
                val num_ty_args =
blanchet@43620
   474
                  length us - the_default 0 (Symtab.lookup sym_tab s)
blanchet@43420
   475
                val (type_us, term_us) =
blanchet@43420
   476
                  chop num_ty_args us |>> append mangled_us
blanchet@43934
   477
                val term_ts = map (do_term [] NONE) term_us
blanchet@43420
   478
                val T =
blanchet@43926
   479
                  if not (null type_us) andalso
blanchet@43926
   480
                     num_type_args thy s' = length type_us then
blanchet@43976
   481
                    (s', map (typ_from_atp ctxt) type_us)
blanchet@43934
   482
                    |> Sign.const_instance thy
blanchet@43626
   483
                  else case opt_T of
blanchet@44023
   484
                    SOME T => map slack_fastype_of term_ts ---> T
blanchet@43626
   485
                  | NONE => HOLogic.typeT
blanchet@43620
   486
                val s' = s' |> unproxify_const
blanchet@43620
   487
              in list_comb (Const (s', T), term_ts @ extra_ts) end
blanchet@43420
   488
          end
blanchet@36901
   489
        | NONE => (* a free or schematic variable *)
blanchet@36901
   490
          let
blanchet@43972
   491
            val ts = map (do_term [] NONE) us @ extra_ts
blanchet@44023
   492
            val T = map slack_fastype_of ts ---> HOLogic.typeT
blanchet@36901
   493
            val t =
blanchet@38987
   494
              case strip_prefix_and_unascii fixed_var_prefix a of
blanchet@36901
   495
                SOME b => Free (b, T)
blanchet@36901
   496
              | NONE =>
blanchet@38987
   497
                case strip_prefix_and_unascii schematic_var_prefix a of
blanchet@43935
   498
                  SOME b => Var ((b, var_index), T)
blanchet@36901
   499
                | NONE =>
blanchet@43936
   500
                  Var ((a |> textual ? repair_variable_name Char.toLower,
blanchet@43936
   501
                        var_index), T)
blanchet@36901
   502
          in list_comb (t, ts) end
blanchet@43934
   503
  in do_term [] end
paulson@21978
   504
blanchet@43976
   505
fun term_from_atom ctxt textual sym_tab pos (u as ATerm (s, _)) =
blanchet@38248
   506
  if String.isPrefix class_prefix s then
blanchet@43976
   507
    add_type_constraint pos (type_constraint_from_term ctxt u)
blanchet@38248
   508
    #> pair @{const True}
blanchet@38248
   509
  else
blanchet@43976
   510
    pair (term_from_atp ctxt textual sym_tab (SOME @{typ bool}) u)
blanchet@36402
   511
blanchet@36553
   512
val combinator_table =
blanchet@40134
   513
  [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def_raw}),
blanchet@40134
   514
   (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def_raw}),
blanchet@40134
   515
   (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def_raw}),
blanchet@40134
   516
   (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def_raw}),
blanchet@40134
   517
   (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def_raw})]
blanchet@36553
   518
blanchet@43626
   519
fun uncombine_term thy =
blanchet@43626
   520
  let
blanchet@43626
   521
    fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
blanchet@43626
   522
      | aux (Abs (s, T, t')) = Abs (s, T, aux t')
blanchet@43626
   523
      | aux (t as Const (x as (s, _))) =
blanchet@43626
   524
        (case AList.lookup (op =) combinator_table s of
blanchet@43626
   525
           SOME thm => thm |> prop_of |> specialize_type thy x
blanchet@43626
   526
                           |> Logic.dest_equals |> snd
blanchet@43626
   527
         | NONE => t)
blanchet@43626
   528
      | aux t = t
blanchet@43626
   529
  in aux end
blanchet@36553
   530
blanchet@38225
   531
(* Update schematic type variables with detected sort constraints. It's not
blanchet@43434
   532
   totally clear whether this code is necessary. *)
blanchet@38248
   533
fun repair_tvar_sorts (t, tvar_tab) =
blanchet@36901
   534
  let
blanchet@38225
   535
    fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
blanchet@38225
   536
      | do_type (TVar (xi, s)) =
blanchet@38225
   537
        TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
blanchet@38225
   538
      | do_type (TFree z) = TFree z
blanchet@38225
   539
    fun do_term (Const (a, T)) = Const (a, do_type T)
blanchet@38225
   540
      | do_term (Free (a, T)) = Free (a, do_type T)
blanchet@38225
   541
      | do_term (Var (xi, T)) = Var (xi, do_type T)
blanchet@38225
   542
      | do_term (t as Bound _) = t
blanchet@38225
   543
      | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)
blanchet@38225
   544
      | do_term (t1 $ t2) = do_term t1 $ do_term t2
blanchet@38225
   545
  in t |> not (Vartab.is_empty tvar_tab) ? do_term end
blanchet@38225
   546
blanchet@39671
   547
fun quantify_over_var quant_of var_s t =
blanchet@39671
   548
  let
blanchet@39671
   549
    val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s)
blanchet@39671
   550
                  |> map Var
blanchet@39671
   551
  in fold_rev quant_of vars t end
blanchet@38225
   552
blanchet@38331
   553
(* Interpret an ATP formula as a HOL term, extracting sort constraints as they
blanchet@38331
   554
   appear in the formula. *)
blanchet@43977
   555
fun raw_prop_from_atp ctxt textual sym_tab phi =
blanchet@38248
   556
  let
blanchet@38248
   557
    fun do_formula pos phi =
blanchet@38225
   558
      case phi of
blanchet@38248
   559
        AQuant (_, [], phi) => do_formula pos phi
blanchet@43397
   560
      | AQuant (q, (s, _) :: xs, phi') =>
blanchet@38248
   561
        do_formula pos (AQuant (q, xs, phi'))
blanchet@43397
   562
        (* FIXME: TFF *)
blanchet@39671
   563
        #>> quantify_over_var (case q of
blanchet@39671
   564
                                 AForall => forall_of
blanchet@39671
   565
                               | AExists => exists_of)
blanchet@43936
   566
                              (s |> textual ? repair_variable_name Char.toLower)
blanchet@38248
   567
      | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
blanchet@38225
   568
      | AConn (c, [phi1, phi2]) =>
blanchet@38248
   569
        do_formula (pos |> c = AImplies ? not) phi1
blanchet@38248
   570
        ##>> do_formula pos phi2
blanchet@38248
   571
        #>> (case c of
blanchet@38248
   572
               AAnd => s_conj
blanchet@38248
   573
             | AOr => s_disj
blanchet@38248
   574
             | AImplies => s_imp
blanchet@38284
   575
             | AIff => s_iff
blanchet@44004
   576
             | ANot => raise Fail "impossible connective")
blanchet@43976
   577
      | AAtom tm => term_from_atom ctxt textual sym_tab pos tm
blanchet@38225
   578
      | _ => raise FORMULA [phi]
blanchet@38248
   579
  in repair_tvar_sorts (do_formula true phi Vartab.empty) end
blanchet@38225
   580
blanchet@43972
   581
fun infer_formula_types ctxt =
wenzelm@39541
   582
  Type.constraint HOLogic.boolT
blanchet@43626
   583
  #> Syntax.check_term
blanchet@43626
   584
         (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
paulson@21978
   585
blanchet@43977
   586
fun prop_from_atp ctxt textual sym_tab =
blanchet@43977
   587
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43977
   588
    raw_prop_from_atp ctxt textual sym_tab
blanchet@44017
   589
    #> textual ? uncombine_term thy #> infer_formula_types ctxt
blanchet@43977
   590
  end
blanchet@43977
   591
blanchet@43934
   592
(**** Translation of TSTP files to Isar proofs ****)
paulson@21978
   593
blanchet@36482
   594
fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
blanchet@36482
   595
  | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
paulson@21978
   596
blanchet@43976
   597
fun decode_line sym_tab (Definition (name, phi1, phi2)) ctxt =
blanchet@36482
   598
    let
wenzelm@43232
   599
      val thy = Proof_Context.theory_of ctxt
blanchet@43977
   600
      (* FIXME: prop or term? *)
blanchet@43977
   601
      val t1 = raw_prop_from_atp ctxt true sym_tab phi1
blanchet@36549
   602
      val vars = snd (strip_comb t1)
blanchet@36482
   603
      val frees = map unvarify_term vars
blanchet@36482
   604
      val unvarify_args = subst_atomic (vars ~~ frees)
blanchet@43977
   605
      val t2 = raw_prop_from_atp ctxt true sym_tab phi2
blanchet@36549
   606
      val (t1, t2) =
blanchet@36549
   607
        HOLogic.eq_const HOLogic.typeT $ t1 $ t2
blanchet@43972
   608
        |> unvarify_args |> uncombine_term thy |> infer_formula_types ctxt
blanchet@36553
   609
        |> HOLogic.dest_eq
blanchet@36482
   610
    in
blanchet@39614
   611
      (Definition (name, t1, t2),
blanchet@36549
   612
       fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt)
blanchet@36482
   613
    end
blanchet@43976
   614
  | decode_line sym_tab (Inference (name, u, deps)) ctxt =
blanchet@43977
   615
    let val t = u |> prop_from_atp ctxt true sym_tab in
blanchet@39614
   616
      (Inference (name, t, deps),
blanchet@36549
   617
       fold Variable.declare_term (OldTerm.term_frees t) ctxt)
blanchet@36482
   618
    end
blanchet@43976
   619
fun decode_lines ctxt sym_tab lines =
blanchet@43976
   620
  fst (fold_map (decode_line sym_tab) lines ctxt)
paulson@21978
   621
blanchet@38281
   622
fun is_same_inference _ (Definition _) = false
blanchet@38281
   623
  | is_same_inference t (Inference (_, t', _)) = t aconv t'
blanchet@36482
   624
blanchet@36482
   625
(* No "real" literals means only type information (tfree_tcs, clsrel, or
blanchet@36482
   626
   clsarity). *)
blanchet@36482
   627
val is_only_type_information = curry (op aconv) HOLogic.true_const
blanchet@36482
   628
blanchet@39619
   629
fun replace_one_dependency (old, new) dep =
blanchet@43809
   630
  if is_same_atp_step dep old then new else [dep]
blanchet@39619
   631
fun replace_dependencies_in_line _ (line as Definition _) = line
blanchet@39619
   632
  | replace_dependencies_in_line p (Inference (name, t, deps)) =
blanchet@39619
   633
    Inference (name, t, fold (union (op =) o replace_one_dependency p) deps [])
paulson@21978
   634
blanchet@40445
   635
(* Discard facts; consolidate adjacent lines that prove the same formula, since
blanchet@38331
   636
   they differ only in type information.*)
blanchet@43518
   637
fun add_line _ _ _ (line as Definition _) lines = line :: lines
blanchet@43518
   638
  | add_line type_sys conjecture_shape fact_names (Inference (name, t, []))
blanchet@43518
   639
             lines =
blanchet@40445
   640
    (* No dependencies: fact, conjecture, or (for Vampire) internal facts or
blanchet@38331
   641
       definitions. *)
blanchet@43518
   642
    if is_fact type_sys fact_names name then
blanchet@40445
   643
      (* Facts are not proof lines. *)
blanchet@36482
   644
      if is_only_type_information t then
blanchet@39619
   645
        map (replace_dependencies_in_line (name, [])) lines
blanchet@36482
   646
      (* Is there a repetition? If so, replace later line by earlier one. *)
blanchet@38281
   647
      else case take_prefix (not o is_same_inference t) lines of
blanchet@39619
   648
        (_, []) => lines (* no repetition of proof line *)
blanchet@39614
   649
      | (pre, Inference (name', _, _) :: post) =>
blanchet@39619
   650
        pre @ map (replace_dependencies_in_line (name', [name])) post
blanchet@40250
   651
      | _ => raise Fail "unexpected inference"
blanchet@39616
   652
    else if is_conjecture conjecture_shape name then
blanchet@43477
   653
      Inference (name, s_not t, []) :: lines
blanchet@36482
   654
    else
blanchet@39619
   655
      map (replace_dependencies_in_line (name, [])) lines
blanchet@43518
   656
  | add_line _ _ _ (Inference (name, t, deps)) lines =
blanchet@36482
   657
    (* Type information will be deleted later; skip repetition test. *)
blanchet@36482
   658
    if is_only_type_information t then
blanchet@39614
   659
      Inference (name, t, deps) :: lines
blanchet@36482
   660
    (* Is there a repetition? If so, replace later line by earlier one. *)
blanchet@38281
   661
    else case take_prefix (not o is_same_inference t) lines of
blanchet@36482
   662
      (* FIXME: Doesn't this code risk conflating proofs involving different
blanchet@38281
   663
         types? *)
blanchet@39614
   664
       (_, []) => Inference (name, t, deps) :: lines
blanchet@39614
   665
     | (pre, Inference (name', t', _) :: post) =>
blanchet@39614
   666
       Inference (name, t', deps) ::
blanchet@39619
   667
       pre @ map (replace_dependencies_in_line (name', [name])) post
blanchet@40250
   668
     | _ => raise Fail "unexpected inference"
paulson@22044
   669
blanchet@36482
   670
(* Recursively delete empty lines (type information) from the proof. *)
blanchet@39614
   671
fun add_nontrivial_line (Inference (name, t, [])) lines =
blanchet@39619
   672
    if is_only_type_information t then delete_dependency name lines
blanchet@39614
   673
    else Inference (name, t, []) :: lines
blanchet@36482
   674
  | add_nontrivial_line line lines = line :: lines
blanchet@39619
   675
and delete_dependency name lines =
blanchet@39619
   676
  fold_rev add_nontrivial_line
blanchet@39619
   677
           (map (replace_dependencies_in_line (name, [])) lines) []
paulson@22470
   678
blanchet@37320
   679
(* ATPs sometimes reuse free variable names in the strangest ways. Removing
blanchet@37320
   680
   offending lines often does the trick. *)
blanchet@36558
   681
fun is_bad_free frees (Free x) = not (member (op =) frees x)
blanchet@36558
   682
  | is_bad_free _ _ = false
paulson@22731
   683
blanchet@43518
   684
fun add_desired_line _ _ _ _ _ (line as Definition (name, _, _)) (j, lines) =
blanchet@39619
   685
    (j, line :: map (replace_dependencies_in_line (name, [])) lines)
blanchet@43518
   686
  | add_desired_line type_sys isar_shrink_factor conjecture_shape fact_names
blanchet@43518
   687
                     frees (Inference (name, t, deps)) (j, lines) =
blanchet@36402
   688
    (j + 1,
blanchet@43518
   689
     if is_fact type_sys fact_names name orelse
blanchet@39616
   690
        is_conjecture conjecture_shape name orelse
blanchet@39619
   691
        (* the last line must be kept *)
blanchet@39619
   692
        j = 0 orelse
blanchet@36568
   693
        (not (is_only_type_information t) andalso
blanchet@36568
   694
         null (Term.add_tvars t []) andalso
blanchet@36568
   695
         not (exists_subterm (is_bad_free frees) t) andalso
blanchet@39619
   696
         length deps >= 2 andalso j mod isar_shrink_factor = 0 andalso
blanchet@39619
   697
         (* kill next to last line, which usually results in a trivial step *)
blanchet@39619
   698
         j <> 1) then
blanchet@39614
   699
       Inference (name, t, deps) :: lines  (* keep line *)
blanchet@36402
   700
     else
blanchet@39619
   701
       map (replace_dependencies_in_line (name, deps)) lines)  (* drop line *)
paulson@21978
   702
blanchet@36482
   703
(** Isar proof construction and manipulation **)
blanchet@36482
   704
blanchet@36482
   705
fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
blanchet@36482
   706
  (union (op =) ls1 ls2, union (op =) ss1 ss2)
blanchet@36402
   707
blanchet@36402
   708
type label = string * int
blanchet@36402
   709
type facts = label list * string list
blanchet@36402
   710
blanchet@39692
   711
datatype isar_qualifier = Show | Then | Moreover | Ultimately
blanchet@36402
   712
blanchet@39692
   713
datatype isar_step =
blanchet@36474
   714
  Fix of (string * typ) list |
blanchet@36482
   715
  Let of term * term |
blanchet@36402
   716
  Assume of label * term |
blanchet@39692
   717
  Have of isar_qualifier list * label * term * byline
blanchet@36402
   718
and byline =
blanchet@36562
   719
  ByMetis of facts |
blanchet@39692
   720
  CaseSplit of isar_step list list * facts
blanchet@36402
   721
blanchet@36572
   722
fun smart_case_split [] facts = ByMetis facts
blanchet@36572
   723
  | smart_case_split proofs facts = CaseSplit (proofs, facts)
blanchet@36572
   724
blanchet@43518
   725
fun add_fact_from_dependency type_sys conjecture_shape facts_offset fact_names
blanchet@43518
   726
                             name =
blanchet@43518
   727
  if is_fact type_sys fact_names name then
blanchet@43518
   728
    apsnd (union (op =)
blanchet@43518
   729
          (map fst (resolve_fact type_sys facts_offset fact_names name)))
blanchet@36471
   730
  else
blanchet@39616
   731
    apfst (insert (op =) (raw_label_for_name conjecture_shape name))
blanchet@36402
   732
blanchet@43518
   733
fun step_for_line _ _ _ _ _ (Definition (_, t1, t2)) = Let (t1, t2)
blanchet@43518
   734
  | step_for_line _ conjecture_shape _ _ _ (Inference (name, t, [])) =
blanchet@39616
   735
    Assume (raw_label_for_name conjecture_shape name, t)
blanchet@43518
   736
  | step_for_line type_sys conjecture_shape facts_offset
blanchet@43518
   737
                  fact_names j (Inference (name, t, deps)) =
blanchet@39616
   738
    Have (if j = 1 then [Show] else [],
blanchet@39671
   739
          raw_label_for_name conjecture_shape name,
blanchet@39671
   740
          fold_rev forall_of (map Var (Term.add_vars t [])) t,
blanchet@43518
   741
          ByMetis (fold (add_fact_from_dependency type_sys conjecture_shape
blanchet@43518
   742
                                                  facts_offset fact_names)
blanchet@39619
   743
                        deps ([], [])))
blanchet@36291
   744
blanchet@39694
   745
fun repair_name "$true" = "c_True"
blanchet@39694
   746
  | repair_name "$false" = "c_False"
blanchet@43841
   747
  | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
blanchet@39694
   748
  | repair_name s =
blanchet@43841
   749
    if is_tptp_equal s orelse
blanchet@43841
   750
       (* seen in Vampire proofs *)
blanchet@43841
   751
       (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
blanchet@43841
   752
      tptp_equal
blanchet@39694
   753
    else
blanchet@39694
   754
      s
blanchet@39694
   755
blanchet@43976
   756
fun isar_proof_from_atp_proof pool ctxt type_sys isar_shrink_factor
blanchet@43874
   757
        conjecture_shape facts_offset fact_names sym_tab params frees
blanchet@43874
   758
        atp_proof =
wenzelm@33318
   759
  let
blanchet@36482
   760
    val lines =
blanchet@43320
   761
      atp_proof
blanchet@43809
   762
      |> clean_up_atp_proof_dependencies
blanchet@39694
   763
      |> nasty_atp_proof pool
blanchet@39694
   764
      |> map_term_names_in_atp_proof repair_name
blanchet@43976
   765
      |> decode_lines ctxt sym_tab
blanchet@43518
   766
      |> rpair [] |-> fold_rev (add_line type_sys conjecture_shape fact_names)
blanchet@36482
   767
      |> rpair [] |-> fold_rev add_nontrivial_line
blanchet@43518
   768
      |> rpair (0, [])
blanchet@43518
   769
      |-> fold_rev (add_desired_line type_sys isar_shrink_factor
blanchet@43518
   770
                                     conjecture_shape fact_names frees)
blanchet@36482
   771
      |> snd
blanchet@36402
   772
  in
blanchet@36901
   773
    (if null params then [] else [Fix params]) @
blanchet@43518
   774
    map2 (step_for_line type_sys conjecture_shape facts_offset fact_names)
blanchet@43412
   775
         (length lines downto 1) lines
blanchet@36402
   776
  end
blanchet@36402
   777
blanchet@36402
   778
(* When redirecting proofs, we keep information about the labels seen so far in
blanchet@36402
   779
   the "backpatches" data structure. The first component indicates which facts
blanchet@36402
   780
   should be associated with forthcoming proof steps. The second component is a
blanchet@37319
   781
   pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
blanchet@37319
   782
   become assumptions and "drop_ls" are the labels that should be dropped in a
blanchet@37319
   783
   case split. *)
blanchet@36402
   784
type backpatches = (label * facts) list * (label list * label list)
blanchet@36402
   785
blanchet@36554
   786
fun used_labels_of_step (Have (_, _, _, by)) =
blanchet@36402
   787
    (case by of
blanchet@36562
   788
       ByMetis (ls, _) => ls
blanchet@36554
   789
     | CaseSplit (proofs, (ls, _)) =>
blanchet@36554
   790
       fold (union (op =) o used_labels_of) proofs ls)
blanchet@36554
   791
  | used_labels_of_step _ = []
blanchet@36554
   792
and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
blanchet@36402
   793
blanchet@36402
   794
fun new_labels_of_step (Fix _) = []
blanchet@36482
   795
  | new_labels_of_step (Let _) = []
blanchet@36402
   796
  | new_labels_of_step (Assume (l, _)) = [l]
blanchet@36402
   797
  | new_labels_of_step (Have (_, l, _, _)) = [l]
blanchet@36402
   798
val new_labels_of = maps new_labels_of_step
blanchet@36402
   799
blanchet@36402
   800
val join_proofs =
blanchet@36402
   801
  let
blanchet@36402
   802
    fun aux _ [] = NONE
blanchet@36402
   803
      | aux proof_tail (proofs as (proof1 :: _)) =
blanchet@36402
   804
        if exists null proofs then
blanchet@36402
   805
          NONE
blanchet@36402
   806
        else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then
blanchet@36402
   807
          aux (hd proof1 :: proof_tail) (map tl proofs)
blanchet@36402
   808
        else case hd proof1 of
blanchet@37498
   809
          Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *)
blanchet@36402
   810
          if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t')
blanchet@36402
   811
                      | _ => false) (tl proofs) andalso
blanchet@36402
   812
             not (exists (member (op =) (maps new_labels_of proofs))
blanchet@36554
   813
                         (used_labels_of proof_tail)) then
blanchet@36402
   814
            SOME (l, t, map rev proofs, proof_tail)
blanchet@36402
   815
          else
blanchet@36402
   816
            NONE
blanchet@36402
   817
        | _ => NONE
blanchet@36402
   818
  in aux [] o map rev end
blanchet@36402
   819
blanchet@36402
   820
fun case_split_qualifiers proofs =
blanchet@36402
   821
  case length proofs of
blanchet@36402
   822
    0 => []
blanchet@36402
   823
  | 1 => [Then]
blanchet@36402
   824
  | _ => [Ultimately]
blanchet@36402
   825
blanchet@39618
   826
fun redirect_proof hyp_ts concl_t proof =
blanchet@36402
   827
  let
blanchet@37321
   828
    (* The first pass outputs those steps that are independent of the negated
blanchet@37321
   829
       conjecture. The second pass flips the proof by contradiction to obtain a
blanchet@37321
   830
       direct proof, introducing case splits when an inference depends on
blanchet@37321
   831
       several facts that depend on the negated conjecture. *)
blanchet@39618
   832
     val concl_l = (conjecture_prefix, length hyp_ts)
blanchet@38286
   833
     fun first_pass ([], contra) = ([], contra)
blanchet@38286
   834
       | first_pass ((step as Fix _) :: proof, contra) =
blanchet@38286
   835
         first_pass (proof, contra) |>> cons step
blanchet@38286
   836
       | first_pass ((step as Let _) :: proof, contra) =
blanchet@38286
   837
         first_pass (proof, contra) |>> cons step
blanchet@39616
   838
       | first_pass ((step as Assume (l as (_, j), _)) :: proof, contra) =
blanchet@39618
   839
         if l = concl_l then first_pass (proof, contra ||> cons step)
blanchet@39618
   840
         else first_pass (proof, contra) |>> cons (Assume (l, nth hyp_ts j))
blanchet@38286
   841
       | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) =
blanchet@39618
   842
         let val step = Have (qs, l, t, ByMetis (ls, ss)) in
blanchet@38286
   843
           if exists (member (op =) (fst contra)) ls then
blanchet@38286
   844
             first_pass (proof, contra |>> cons l ||> cons step)
blanchet@38286
   845
           else
blanchet@38286
   846
             first_pass (proof, contra) |>> cons step
blanchet@38286
   847
         end
blanchet@38286
   848
       | first_pass _ = raise Fail "malformed proof"
blanchet@36402
   849
    val (proof_top, (contra_ls, contra_proof)) =
blanchet@39618
   850
      first_pass (proof, ([concl_l], []))
blanchet@36402
   851
    val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst
blanchet@36402
   852
    fun backpatch_labels patches ls =
blanchet@36402
   853
      fold merge_fact_sets (map (backpatch_label patches) ls) ([], [])
blanchet@36402
   854
    fun second_pass end_qs ([], assums, patches) =
blanchet@37321
   855
        ([Have (end_qs, no_label, concl_t,
blanchet@36562
   856
                ByMetis (backpatch_labels patches (map snd assums)))], patches)
blanchet@36402
   857
      | second_pass end_qs (Assume (l, t) :: proof, assums, patches) =
blanchet@36402
   858
        second_pass end_qs (proof, (t, l) :: assums, patches)
blanchet@36562
   859
      | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums,
blanchet@36402
   860
                            patches) =
blanchet@39619
   861
        (if member (op =) (snd (snd patches)) l andalso
blanchet@39619
   862
            not (member (op =) (fst (snd patches)) l) andalso
blanchet@39619
   863
            not (AList.defined (op =) (fst patches) l) then
blanchet@39619
   864
           second_pass end_qs (proof, assums, patches ||> apsnd (append ls))
blanchet@39619
   865
         else case List.partition (member (op =) contra_ls) ls of
blanchet@39619
   866
           ([contra_l], co_ls) =>
blanchet@39619
   867
           if member (op =) qs Show then
blanchet@39619
   868
             second_pass end_qs (proof, assums,
blanchet@39619
   869
                                 patches |>> cons (contra_l, (co_ls, ss)))
blanchet@39619
   870
           else
blanchet@39619
   871
             second_pass end_qs
blanchet@39619
   872
                         (proof, assums,
blanchet@39619
   873
                          patches |>> cons (contra_l, (l :: co_ls, ss)))
blanchet@39619
   874
             |>> cons (if member (op =) (fst (snd patches)) l then
blanchet@43477
   875
                         Assume (l, s_not t)
blanchet@39619
   876
                       else
blanchet@43477
   877
                         Have (qs, l, s_not t,
blanchet@39619
   878
                               ByMetis (backpatch_label patches l)))
blanchet@39619
   879
         | (contra_ls as _ :: _, co_ls) =>
blanchet@39619
   880
           let
blanchet@39619
   881
             val proofs =
blanchet@39619
   882
               map_filter
blanchet@39619
   883
                   (fn l =>
blanchet@39619
   884
                       if l = concl_l then
blanchet@39619
   885
                         NONE
blanchet@39619
   886
                       else
blanchet@39619
   887
                         let
blanchet@39619
   888
                           val drop_ls = filter (curry (op <>) l) contra_ls
blanchet@39619
   889
                         in
blanchet@39619
   890
                           second_pass []
blanchet@39619
   891
                               (proof, assums,
blanchet@39619
   892
                                patches ||> apfst (insert (op =) l)
blanchet@39619
   893
                                        ||> apsnd (union (op =) drop_ls))
blanchet@39619
   894
                           |> fst |> SOME
blanchet@39619
   895
                         end) contra_ls
blanchet@39619
   896
             val (assumes, facts) =
blanchet@39619
   897
               if member (op =) (fst (snd patches)) l then
blanchet@43477
   898
                 ([Assume (l, s_not t)], (l :: co_ls, ss))
blanchet@39619
   899
               else
blanchet@39619
   900
                 ([], (co_ls, ss))
blanchet@39619
   901
           in
blanchet@39619
   902
             (case join_proofs proofs of
blanchet@39619
   903
                SOME (l, t, proofs, proof_tail) =>
blanchet@39619
   904
                Have (case_split_qualifiers proofs @
blanchet@39619
   905
                      (if null proof_tail then end_qs else []), l, t,
blanchet@39619
   906
                      smart_case_split proofs facts) :: proof_tail
blanchet@39619
   907
              | NONE =>
blanchet@39619
   908
                [Have (case_split_qualifiers proofs @ end_qs, no_label,
blanchet@39619
   909
                       concl_t, smart_case_split proofs facts)],
blanchet@39619
   910
              patches)
blanchet@39619
   911
             |>> append assumes
blanchet@39619
   912
           end
blanchet@39619
   913
         | _ => raise Fail "malformed proof")
blanchet@36402
   914
       | second_pass _ _ = raise Fail "malformed proof"
blanchet@36482
   915
    val proof_bottom =
blanchet@36482
   916
      second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
blanchet@36402
   917
  in proof_top @ proof_bottom end
blanchet@36402
   918
blanchet@38715
   919
(* FIXME: Still needed? Probably not. *)
blanchet@36402
   920
val kill_duplicate_assumptions_in_proof =
blanchet@36402
   921
  let
blanchet@36402
   922
    fun relabel_facts subst =
blanchet@36402
   923
      apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
blanchet@36487
   924
    fun do_step (step as Assume (l, t)) (proof, subst, assums) =
blanchet@36402
   925
        (case AList.lookup (op aconv) assums t of
blanchet@36960
   926
           SOME l' => (proof, (l, l') :: subst, assums)
blanchet@36487
   927
         | NONE => (step :: proof, subst, (t, l) :: assums))
blanchet@36402
   928
      | do_step (Have (qs, l, t, by)) (proof, subst, assums) =
blanchet@36402
   929
        (Have (qs, l, t,
blanchet@36402
   930
               case by of
blanchet@36562
   931
                 ByMetis facts => ByMetis (relabel_facts subst facts)
blanchet@36402
   932
               | CaseSplit (proofs, facts) =>
blanchet@36402
   933
                 CaseSplit (map do_proof proofs, relabel_facts subst facts)) ::
blanchet@36402
   934
         proof, subst, assums)
blanchet@36487
   935
      | do_step step (proof, subst, assums) = (step :: proof, subst, assums)
blanchet@36402
   936
    and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev
blanchet@36402
   937
  in do_proof end
blanchet@36402
   938
blanchet@36402
   939
val then_chain_proof =
blanchet@36402
   940
  let
blanchet@36402
   941
    fun aux _ [] = []
blanchet@36487
   942
      | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof
blanchet@36402
   943
      | aux l' (Have (qs, l, t, by) :: proof) =
blanchet@36402
   944
        (case by of
blanchet@36562
   945
           ByMetis (ls, ss) =>
blanchet@36402
   946
           Have (if member (op =) ls l' then
blanchet@36402
   947
                   (Then :: qs, l, t,
blanchet@36562
   948
                    ByMetis (filter_out (curry (op =) l') ls, ss))
blanchet@36402
   949
                 else
blanchet@36562
   950
                   (qs, l, t, ByMetis (ls, ss)))
blanchet@36402
   951
         | CaseSplit (proofs, facts) =>
blanchet@36402
   952
           Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) ::
blanchet@36402
   953
        aux l proof
blanchet@36487
   954
      | aux _ (step :: proof) = step :: aux no_label proof
blanchet@36402
   955
  in aux no_label end
blanchet@36402
   956
blanchet@36402
   957
fun kill_useless_labels_in_proof proof =
blanchet@36402
   958
  let
blanchet@36554
   959
    val used_ls = used_labels_of proof
blanchet@36402
   960
    fun do_label l = if member (op =) used_ls l then l else no_label
blanchet@36554
   961
    fun do_step (Assume (l, t)) = Assume (do_label l, t)
blanchet@36554
   962
      | do_step (Have (qs, l, t, by)) =
blanchet@36402
   963
        Have (qs, do_label l, t,
blanchet@36402
   964
              case by of
blanchet@36402
   965
                CaseSplit (proofs, facts) =>
blanchet@36554
   966
                CaseSplit (map (map do_step) proofs, facts)
blanchet@36402
   967
              | _ => by)
blanchet@36554
   968
      | do_step step = step
blanchet@36554
   969
  in map do_step proof end
blanchet@36402
   970
blanchet@36402
   971
fun prefix_for_depth n = replicate_string (n + 1)
blanchet@36402
   972
blanchet@36402
   973
val relabel_proof =
blanchet@36402
   974
  let
blanchet@36402
   975
    fun aux _ _ _ [] = []
blanchet@36402
   976
      | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) =
blanchet@36402
   977
        if l = no_label then
blanchet@36402
   978
          Assume (l, t) :: aux subst depth (next_assum, next_fact) proof
blanchet@36402
   979
        else
blanchet@36402
   980
          let val l' = (prefix_for_depth depth assum_prefix, next_assum) in
blanchet@36402
   981
            Assume (l', t) ::
blanchet@36402
   982
            aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof
blanchet@36402
   983
          end
blanchet@36402
   984
      | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) =
blanchet@36402
   985
        let
blanchet@36402
   986
          val (l', subst, next_fact) =
blanchet@36402
   987
            if l = no_label then
blanchet@36402
   988
              (l, subst, next_fact)
blanchet@36402
   989
            else
blanchet@36402
   990
              let
blanchet@43051
   991
                val l' = (prefix_for_depth depth have_prefix, next_fact)
blanchet@36402
   992
              in (l', (l, l') :: subst, next_fact + 1) end
blanchet@36568
   993
          val relabel_facts =
blanchet@39616
   994
            apfst (maps (the_list o AList.lookup (op =) subst))
blanchet@36402
   995
          val by =
blanchet@36402
   996
            case by of
blanchet@36562
   997
              ByMetis facts => ByMetis (relabel_facts facts)
blanchet@36402
   998
            | CaseSplit (proofs, facts) =>
blanchet@36402
   999
              CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs,
blanchet@36402
  1000
                         relabel_facts facts)
blanchet@36402
  1001
        in
blanchet@36402
  1002
          Have (qs, l', t, by) ::
blanchet@36402
  1003
          aux subst depth (next_assum, next_fact) proof
blanchet@36402
  1004
        end
blanchet@36487
  1005
      | aux subst depth nextp (step :: proof) =
blanchet@36487
  1006
        step :: aux subst depth nextp proof
blanchet@36402
  1007
  in aux [] 0 (1, 1) end
blanchet@36402
  1008
blanchet@43750
  1009
fun string_for_proof ctxt0 full_types i n =
blanchet@36402
  1010
  let
blanchet@43626
  1011
    val ctxt =
blanchet@43626
  1012
      ctxt0 |> Config.put show_free_types false
blanchet@43626
  1013
            |> Config.put show_types true
blanchet@43626
  1014
            |> Config.put show_sorts true
blanchet@37316
  1015
    fun fix_print_mode f x =
wenzelm@39393
  1016
      Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
wenzelm@39393
  1017
                               (print_mode_value ())) f x
blanchet@36402
  1018
    fun do_indent ind = replicate_string (ind * indent_size) " "
blanchet@36474
  1019
    fun do_free (s, T) =
blanchet@36474
  1020
      maybe_quote s ^ " :: " ^
blanchet@36474
  1021
      maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
blanchet@36568
  1022
    fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
blanchet@36402
  1023
    fun do_have qs =
blanchet@36402
  1024
      (if member (op =) qs Moreover then "moreover " else "") ^
blanchet@36402
  1025
      (if member (op =) qs Ultimately then "ultimately " else "") ^
blanchet@36402
  1026
      (if member (op =) qs Then then
blanchet@36402
  1027
         if member (op =) qs Show then "thus" else "hence"
blanchet@36402
  1028
       else
blanchet@36402
  1029
         if member (op =) qs Show then "show" else "have")
blanchet@36474
  1030
    val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
blanchet@44009
  1031
    val reconstructor = if full_types then MetisX else Metis
blanchet@36568
  1032
    fun do_facts (ls, ss) =
blanchet@43874
  1033
      reconstructor_command reconstructor 1 1
blanchet@43874
  1034
          (ls |> sort_distinct (prod_ord string_ord int_ord),
blanchet@43874
  1035
           ss |> sort_distinct string_ord)
blanchet@36474
  1036
    and do_step ind (Fix xs) =
blanchet@36474
  1037
        do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
blanchet@36482
  1038
      | do_step ind (Let (t1, t2)) =
blanchet@36482
  1039
        do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
blanchet@36402
  1040
      | do_step ind (Assume (l, t)) =
blanchet@36402
  1041
        do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
blanchet@36562
  1042
      | do_step ind (Have (qs, l, t, ByMetis facts)) =
blanchet@36402
  1043
        do_indent ind ^ do_have qs ^ " " ^
blanchet@36475
  1044
        do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n"
blanchet@36402
  1045
      | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) =
blanchet@36402
  1046
        space_implode (do_indent ind ^ "moreover\n")
blanchet@36402
  1047
                      (map (do_block ind) proofs) ^
blanchet@36475
  1048
        do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^
blanchet@36474
  1049
        do_facts facts ^ "\n"
blanchet@36402
  1050
    and do_steps prefix suffix ind steps =
blanchet@36402
  1051
      let val s = implode (map (do_step ind) steps) in
blanchet@36402
  1052
        replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
blanchet@36402
  1053
        String.extract (s, ind * indent_size,
blanchet@36402
  1054
                        SOME (size s - ind * indent_size - 1)) ^
blanchet@36402
  1055
        suffix ^ "\n"
blanchet@36402
  1056
      end
blanchet@36402
  1057
    and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
blanchet@36562
  1058
    (* One-step proofs are pointless; better use the Metis one-liner
blanchet@36562
  1059
       directly. *)
blanchet@36562
  1060
    and do_proof [Have (_, _, _, ByMetis _)] = ""
blanchet@36562
  1061
      | do_proof proof =
blanchet@36476
  1062
        (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
blanchet@39692
  1063
        do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^
blanchet@39692
  1064
        (if n <> 1 then "next" else "qed")
blanchet@36484
  1065
  in do_proof end
blanchet@36402
  1066
blanchet@43903
  1067
fun isar_proof_text ctxt isar_proof_requested
blanchet@43874
  1068
        (debug, full_types, isar_shrink_factor, type_sys, pool,
blanchet@43878
  1069
         conjecture_shape, facts_offset, fact_names, sym_tab, atp_proof, goal)
blanchet@43878
  1070
        (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
blanchet@36402
  1071
  let
blanchet@43903
  1072
    val isar_shrink_factor =
blanchet@43903
  1073
      (if isar_proof_requested then 1 else 2) * isar_shrink_factor
blanchet@43878
  1074
    val (params, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal
blanchet@36901
  1075
    val frees = fold Term.add_frees (concl_t :: hyp_ts) []
blanchet@43874
  1076
    val one_line_proof = one_line_proof_text one_line_params
blanchet@36283
  1077
    fun isar_proof_for () =
blanchet@43874
  1078
      case atp_proof
blanchet@43976
  1079
           |> isar_proof_from_atp_proof pool ctxt type_sys isar_shrink_factor
blanchet@43976
  1080
                  conjecture_shape facts_offset fact_names sym_tab params frees
blanchet@39618
  1081
           |> redirect_proof hyp_ts concl_t
blanchet@36402
  1082
           |> kill_duplicate_assumptions_in_proof
blanchet@36402
  1083
           |> then_chain_proof
blanchet@36402
  1084
           |> kill_useless_labels_in_proof
blanchet@36402
  1085
           |> relabel_proof
blanchet@43878
  1086
           |> string_for_proof ctxt full_types subgoal subgoal_count of
blanchet@43903
  1087
        "" =>
blanchet@43903
  1088
        if isar_proof_requested then
blanchet@43903
  1089
          "\nNo structured proof available (proof too short)."
blanchet@43903
  1090
        else
blanchet@43903
  1091
          ""
blanchet@43903
  1092
      | proof =>
blanchet@43903
  1093
        "\n\n" ^ (if isar_proof_requested then "Structured proof"
blanchet@43903
  1094
                  else "Perhaps this will work") ^
blanchet@43903
  1095
        ":\n" ^ Markup.markup Markup.sendback proof
blanchet@35868
  1096
    val isar_proof =
blanchet@36402
  1097
      if debug then
blanchet@36283
  1098
        isar_proof_for ()
blanchet@36283
  1099
      else
blanchet@43903
  1100
        case try isar_proof_for () of
blanchet@43903
  1101
          SOME s => s
blanchet@43903
  1102
        | NONE => if isar_proof_requested then
blanchet@43903
  1103
                    "\nWarning: The Isar proof construction failed."
blanchet@43903
  1104
                  else
blanchet@43903
  1105
                    ""
blanchet@43874
  1106
  in one_line_proof ^ isar_proof end
immler@31038
  1107
blanchet@43874
  1108
fun proof_text ctxt isar_proof isar_params
blanchet@43874
  1109
               (one_line_params as (preplay, _, _, _, _, _)) =
blanchet@44007
  1110
  (if case preplay of Failed_to_Play _ => true | _ => isar_proof then
blanchet@43903
  1111
     isar_proof_text ctxt isar_proof isar_params
blanchet@43874
  1112
   else
blanchet@43874
  1113
     one_line_proof_text) one_line_params
blanchet@36223
  1114
immler@31038
  1115
end;