src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
author blanchet
Fri, 15 Feb 2013 16:53:39 +0100
changeset 51165 58f8716b04ee
parent 51164 eba05aaa8612
child 51178 06689dbfe072
permissions -rw-r--r--
tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     3
    Author:     Steffen Juilf Smolka, TU Muenchen
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     4
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
     5
Isar proof reconstruction from ATP proofs.
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     6
*)
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     7
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     8
signature SLEDGEHAMMER_PROOF_RECONSTRUCT =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     9
sig
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    10
  type 'a proof = 'a ATP_Proof.proof
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    11
  type stature = ATP_Problem_Generate.stature
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    12
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    13
  datatype reconstructor =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    14
    Metis of string * string |
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    15
    SMT
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    16
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    17
  datatype play =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    18
    Played of reconstructor * Time.time |
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    19
    Trust_Playable of reconstructor * Time.time option |
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    20
    Failed_to_Play of reconstructor
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    21
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    22
  type minimize_command = string list -> string
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    23
  type one_line_params =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    24
    play * string * (string * stature) list * minimize_command * int * int
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    25
  type isar_params =
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50450
diff changeset
    26
    bool * bool * Time.time option * real * string Symtab.table
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
    27
    * (string * stature) list vector * int Symtab.table * string proof * thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    28
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    29
  val smtN : string
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    30
  val string_for_reconstructor : reconstructor -> string
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    31
  val lam_trans_from_atp_proof : string proof -> string -> string
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    32
  val is_typed_helper_used_in_atp_proof : string proof -> bool
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    33
  val used_facts_in_atp_proof :
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    34
    Proof.context -> (string * stature) list vector -> string proof ->
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    35
    (string * stature) list
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    36
  val used_facts_in_unsound_atp_proof :
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    37
    Proof.context -> (string * stature) list vector -> 'a proof ->
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    38
    string list option
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    39
  val one_line_proof_text : int -> one_line_params -> string
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    40
  val isar_proof_text :
49913
blanchet
parents: 49883
diff changeset
    41
    Proof.context -> bool -> isar_params -> one_line_params -> string
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    42
  val proof_text :
49913
blanchet
parents: 49883
diff changeset
    43
    Proof.context -> bool -> isar_params -> int -> one_line_params -> string
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    44
end;
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    45
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    46
structure Sledgehammer_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    47
struct
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    48
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    49
open ATP_Util
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    50
open ATP_Problem
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    51
open ATP_Proof
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    52
open ATP_Problem_Generate
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    53
open ATP_Proof_Reconstruct
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
    54
open Sledgehammer_Util
50264
a9ec48b98734 renamed sledgehammer_isar_reconstruct to sledgehammer_proof
smolkas
parents: 50262
diff changeset
    55
open Sledgehammer_Proof
50258
1c708d7728c7 put annotate in own structure
smolkas
parents: 50257
diff changeset
    56
open Sledgehammer_Annotate
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51129
diff changeset
    57
open Sledgehammer_Compress
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    58
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    59
structure String_Redirect = ATP_Proof_Redirect(
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    60
  type key = step_name
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    61
  val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s')
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    62
  val string_of = fst)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    63
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    64
open String_Redirect
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    65
49916
412346127bfa fixed theorem lookup code in Isar proof reconstruction
blanchet
parents: 49915
diff changeset
    66
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    67
(** reconstructors **)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    68
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    69
datatype reconstructor =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    70
  Metis of string * string |
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    71
  SMT
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    72
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    73
datatype play =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    74
  Played of reconstructor * Time.time |
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    75
  Trust_Playable of reconstructor * Time.time option |
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    76
  Failed_to_Play of reconstructor
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    77
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    78
val smtN = "smt"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    79
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    80
fun string_for_reconstructor (Metis (type_enc, lam_trans)) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    81
    metis_call type_enc lam_trans
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    82
  | string_for_reconstructor SMT = smtN
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    83
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    84
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    85
(** fact extraction from ATP proofs **)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    86
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    87
fun find_first_in_list_vector vec key =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    88
  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    89
                 | (_, value) => value) NONE vec
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    90
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    91
val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    92
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    93
fun resolve_one_named_fact fact_names s =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    94
  case try (unprefix fact_prefix) s of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    95
    SOME s' =>
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    96
    let val s' = s' |> unprefix_fact_number |> unascii_of in
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    97
      s' |> find_first_in_list_vector fact_names |> Option.map (pair s')
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    98
    end
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    99
  | NONE => NONE
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   100
fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   101
fun is_fact fact_names = not o null o resolve_fact fact_names
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   102
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   103
fun resolve_one_named_conjecture s =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   104
  case try (unprefix conjecture_prefix) s of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   105
    SOME s' => Int.fromString s'
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   106
  | NONE => NONE
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   107
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   108
val resolve_conjecture = map_filter resolve_one_named_conjecture
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   109
val is_conjecture = not o null o resolve_conjecture
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   110
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   111
val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   112
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   113
(* overapproximation (good enough) *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   114
fun is_lam_lifted s =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   115
  String.isPrefix fact_prefix s andalso
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   116
  String.isSubstring ascii_of_lam_fact_prefix s
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   117
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   118
val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   119
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   120
fun is_axiom_used_in_proof pred =
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   121
  exists (fn Inference_Step ((_, ss), _, _, _, []) => exists pred ss
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   122
           | _ => false)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   123
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   124
fun lam_trans_from_atp_proof atp_proof default =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   125
  case (is_axiom_used_in_proof is_combinator_def atp_proof,
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   126
        is_axiom_used_in_proof is_lam_lifted atp_proof) of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   127
    (false, false) => default
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   128
  | (false, true) => liftingN
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   129
(*  | (true, true) => combs_and_liftingN -- not supported by "metis" *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   130
  | (true, _) => combsN
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   131
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   132
val is_typed_helper_name =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   133
  String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
51031
63d71b247323 more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
blanchet
parents: 51026
diff changeset
   134
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   135
fun is_typed_helper_used_in_atp_proof atp_proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   136
  is_axiom_used_in_proof is_typed_helper_name atp_proof
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   137
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   138
fun add_non_rec_defs fact_names accum =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   139
  Vector.foldl (fn (facts, facts') =>
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   140
      union (op =) (filter (fn (_, (_, status)) => status = Non_Rec_Def) facts)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   141
            facts')
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   142
    accum fact_names
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   143
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   144
val isa_ext = Thm.get_name_hint @{thm ext}
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   145
val isa_short_ext = Long_Name.base_name isa_ext
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   146
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   147
fun ext_name ctxt =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   148
  if Thm.eq_thm_prop (@{thm ext},
51026
48e82e199df1 tuned indent
blanchet
parents: 50924
diff changeset
   149
       singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   150
    isa_short_ext
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   151
  else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   152
    isa_ext
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   153
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   154
val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   155
val leo2_unfold_def_rule = "unfold_def"
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   156
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   157
fun add_fact ctxt fact_names (Inference_Step ((_, ss), _, _, rule, deps)) =
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   158
    (if rule = leo2_extcnf_equal_neg_rule then
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   159
       insert (op =) (ext_name ctxt, (Global, General))
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   160
     else if rule = leo2_unfold_def_rule then
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   161
       (* LEO 1.3.3 does not record definitions properly, leading to missing
51026
48e82e199df1 tuned indent
blanchet
parents: 50924
diff changeset
   162
          dependencies in the TSTP proof. Remove the next line once this is
48e82e199df1 tuned indent
blanchet
parents: 50924
diff changeset
   163
          fixed. *)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   164
       add_non_rec_defs fact_names
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   165
     else if rule = satallax_coreN then
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   166
       (fn [] =>
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   167
           (* Satallax doesn't include definitions in its unsatisfiable cores,
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   168
              so we assume the worst and include them all here. *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   169
           [(ext_name ctxt, (Global, General))] |> add_non_rec_defs fact_names
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   170
         | facts => facts)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   171
     else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   172
       I)
51026
48e82e199df1 tuned indent
blanchet
parents: 50924
diff changeset
   173
    #> (if null deps then union (op =) (resolve_fact fact_names ss) else I)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   174
  | add_fact _ _ _ = I
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   175
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   176
fun used_facts_in_atp_proof ctxt fact_names atp_proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   177
  if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   178
  else fold (add_fact ctxt fact_names) atp_proof []
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   179
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   180
fun used_facts_in_unsound_atp_proof _ _ [] = NONE
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   181
  | used_facts_in_unsound_atp_proof ctxt fact_names atp_proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   182
    let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   183
      if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   184
         not (is_axiom_used_in_proof (is_conjecture o single) atp_proof) then
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   185
        SOME (map fst used_facts)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   186
      else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   187
        NONE
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   188
    end
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   189
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   190
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   191
(** one-liner reconstructor proofs **)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   192
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   193
fun show_time NONE = ""
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   194
  | show_time (SOME ext_time) = " (" ^ string_from_ext_time ext_time ^ ")"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   195
49986
90e7be285b49 took out "using only ..." comments in Sledgehammer generated metis/smt calls, until these can be generated soundly
blanchet
parents: 49983
diff changeset
   196
(* FIXME: Various bugs, esp. with "unfolding"
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   197
fun unusing_chained_facts _ 0 = ""
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   198
  | unusing_chained_facts used_chaineds num_chained =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   199
    if length used_chaineds = num_chained then ""
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   200
    else if null used_chaineds then "(* using no facts *) "
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   201
    else "(* using only " ^ space_implode " " used_chaineds ^ " *) "
49986
90e7be285b49 took out "using only ..." comments in Sledgehammer generated metis/smt calls, until these can be generated soundly
blanchet
parents: 49983
diff changeset
   202
*)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   203
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   204
fun apply_on_subgoal _ 1 = "by "
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   205
  | apply_on_subgoal 1 _ = "apply "
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   206
  | apply_on_subgoal i n =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   207
    "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   208
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   209
fun using_labels [] = ""
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   210
  | using_labels ls =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   211
    "using " ^ space_implode " " (map string_for_label ls) ^ " "
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   212
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   213
fun command_call name [] =
50239
fb579401dc26 tuned signature;
wenzelm
parents: 50218
diff changeset
   214
    name |> not (Symbol_Pos.is_identifier name) ? enclose "(" ")"
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   215
  | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   216
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   217
fun reconstructor_command reconstr i n used_chaineds num_chained (ls, ss) =
49986
90e7be285b49 took out "using only ..." comments in Sledgehammer generated metis/smt calls, until these can be generated soundly
blanchet
parents: 49983
diff changeset
   218
  (* unusing_chained_facts used_chaineds num_chained ^ *)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   219
  using_labels ls ^ apply_on_subgoal i n ^
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   220
  command_call (string_for_reconstructor reconstr) ss
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   221
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   222
fun try_command_line banner time command =
50450
358b6020f8b6 generalized notion of active area, where sendback is just one application;
wenzelm
parents: 50410
diff changeset
   223
  banner ^ ": " ^ Active.sendback_markup command ^ show_time time ^ "."
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   224
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   225
fun minimize_line _ [] = ""
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   226
  | minimize_line minimize_command ss =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   227
    case minimize_command ss of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   228
      "" => ""
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   229
    | command =>
50450
358b6020f8b6 generalized notion of active area, where sendback is just one application;
wenzelm
parents: 50410
diff changeset
   230
      "\nTo minimize: " ^ Active.sendback_markup command ^ "."
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   231
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   232
fun split_used_facts facts =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   233
  facts |> List.partition (fn (_, (sc, _)) => sc = Chained)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   234
        |> pairself (sort_distinct (string_ord o pairself fst))
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   235
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   236
type minimize_command = string list -> string
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   237
type one_line_params =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   238
  play * string * (string * stature) list * minimize_command * int * int
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   239
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   240
fun one_line_proof_text num_chained
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   241
        (preplay, banner, used_facts, minimize_command, subgoal,
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   242
         subgoal_count) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   243
  let
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   244
    val (chained, extra) = split_used_facts used_facts
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   245
    val (failed, reconstr, ext_time) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   246
      case preplay of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   247
        Played (reconstr, time) => (false, reconstr, (SOME (false, time)))
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   248
      | Trust_Playable (reconstr, time) =>
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   249
        (false, reconstr,
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   250
         case time of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   251
           NONE => NONE
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   252
         | SOME time =>
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   253
           if time = Time.zeroTime then NONE else SOME (true, time))
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   254
      | Failed_to_Play reconstr => (true, reconstr, NONE)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   255
    val try_line =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   256
      ([], map fst extra)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   257
      |> reconstructor_command reconstr subgoal subgoal_count (map fst chained)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   258
                               num_chained
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   259
      |> (if failed then
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   260
            enclose "One-line proof reconstruction failed: "
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   261
                     ".\n(Invoking \"sledgehammer\" with \"[strict]\" might \
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   262
                     \solve this.)"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   263
          else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   264
            try_command_line banner ext_time)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   265
  in try_line ^ minimize_line minimize_command (map fst (extra @ chained)) end
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   266
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   267
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   268
(** Isar proof construction and manipulation **)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   269
50017
d9c1b11a78d2 avoid name clashes
blanchet
parents: 50016
diff changeset
   270
val assume_prefix = "a"
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   271
val have_prefix = "f"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   272
val raw_prefix = "x"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   273
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   274
fun raw_label_for_name (num, ss) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   275
  case resolve_conjecture ss of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   276
    [j] => (conjecture_prefix, j)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   277
  | _ => (raw_prefix ^ ascii_of num, 0)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   278
50005
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   279
fun label_of_clause [name] = raw_label_for_name name
50262
6dc80eead659 fixed problem with fact names
smolkas
parents: 50261
diff changeset
   280
  | label_of_clause c = (space_implode "___" (map (fst o raw_label_for_name) c), 0)
50005
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   281
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   282
fun add_fact_from_dependencies fact_names (names as [(_, ss)]) =
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   283
    if is_fact fact_names ss then
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   284
      apsnd (union (op =) (map fst (resolve_fact fact_names ss)))
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   285
    else
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   286
      apfst (insert (op =) (label_of_clause names))
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   287
  | add_fact_from_dependencies fact_names names =
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
   288
    apfst (insert (op =) (label_of_clause names))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   289
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   290
fun repair_name "$true" = "c_True"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   291
  | repair_name "$false" = "c_False"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   292
  | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   293
  | repair_name s =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   294
    if is_tptp_equal s orelse
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   295
       (* seen in Vampire proofs *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   296
       (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   297
      tptp_equal
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   298
    else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   299
      s
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   300
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   301
fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   302
  | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   303
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   304
fun infer_formula_types ctxt =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   305
  Type.constraint HOLogic.boolT
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   306
  #> Syntax.check_term
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   307
         (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   308
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   309
val combinator_table =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   310
  [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   311
   (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   312
   (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   313
   (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   314
   (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   315
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   316
fun uncombine_term thy =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   317
  let
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   318
    fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   319
      | aux (Abs (s, T, t')) = Abs (s, T, aux t')
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   320
      | aux (t as Const (x as (s, _))) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   321
        (case AList.lookup (op =) combinator_table s of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   322
           SOME thm => thm |> prop_of |> specialize_type thy x
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   323
                           |> Logic.dest_equals |> snd
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   324
         | NONE => t)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   325
      | aux t = t
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   326
  in aux end
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   327
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   328
fun decode_line sym_tab (Definition_Step (name, phi1, phi2)) ctxt =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   329
    let
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   330
      val thy = Proof_Context.theory_of ctxt
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   331
      val t1 = prop_from_atp ctxt true sym_tab phi1
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   332
      val vars = snd (strip_comb t1)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   333
      val frees = map unvarify_term vars
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   334
      val unvarify_args = subst_atomic (vars ~~ frees)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   335
      val t2 = prop_from_atp ctxt true sym_tab phi2
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   336
      val (t1, t2) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   337
        HOLogic.eq_const HOLogic.typeT $ t1 $ t2
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   338
        |> unvarify_args |> uncombine_term thy |> infer_formula_types ctxt
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   339
        |> HOLogic.dest_eq
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   340
    in
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   341
      (Definition_Step (name, t1, t2),
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   342
       fold Variable.declare_term (maps Misc_Legacy.term_frees [t1, t2]) ctxt)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   343
    end
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   344
  | decode_line sym_tab (Inference_Step (name, role, u, rule, deps)) ctxt =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   345
    let
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   346
      val thy = Proof_Context.theory_of ctxt
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   347
      val t = u |> prop_from_atp ctxt true sym_tab
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   348
                |> uncombine_term thy |> infer_formula_types ctxt
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   349
    in
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   350
      (Inference_Step (name, role, t, rule, deps),
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   351
       fold Variable.declare_term (Misc_Legacy.term_frees t) ctxt)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   352
    end
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   353
fun decode_lines ctxt sym_tab lines =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   354
  fst (fold_map (decode_line sym_tab) lines ctxt)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   355
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   356
fun replace_one_dependency (old, new) dep =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   357
  if is_same_atp_step dep old then new else [dep]
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   358
fun replace_dependencies_in_line _ (line as Definition_Step _) = line
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   359
  | replace_dependencies_in_line p
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   360
        (Inference_Step (name, role, t, rule, deps)) =
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   361
    Inference_Step (name, role, t, rule,
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   362
                    fold (union (op =) o replace_one_dependency p) deps [])
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   363
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   364
(* No "real" literals means only type information (tfree_tcs, clsrel, or
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   365
   clsarity). *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   366
fun is_only_type_information t = t aconv @{term True}
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   367
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   368
fun s_maybe_not role = role <> Conjecture ? s_not
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   369
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   370
fun is_same_inference _ (Definition_Step _) = false
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   371
  | is_same_inference (role, t) (Inference_Step (_, role', t', _, _)) =
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   372
    s_maybe_not role t aconv s_maybe_not role' t'
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   373
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   374
(* Discard facts; consolidate adjacent lines that prove the same formula, since
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   375
   they differ only in type information.*)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   376
fun add_line _ (line as Definition_Step _) lines = line :: lines
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   377
  | add_line fact_names (Inference_Step (name as (_, ss), role, t, rule, []))
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   378
             lines =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   379
    (* No dependencies: fact, conjecture, or (for Vampire) internal facts or
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   380
       definitions. *)
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   381
    if is_conjecture ss then
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   382
      Inference_Step (name, role, t, rule, []) :: lines
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   383
    else if is_fact fact_names ss then
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   384
      (* Facts are not proof lines. *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   385
      if is_only_type_information t then
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   386
        map (replace_dependencies_in_line (name, [])) lines
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   387
      else
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   388
        lines
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   389
    else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   390
      map (replace_dependencies_in_line (name, [])) lines
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   391
  | add_line _ (line as Inference_Step (name, role, t, rule, deps)) lines =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   392
    (* Type information will be deleted later; skip repetition test. *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   393
    if is_only_type_information t then
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   394
      line :: lines
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   395
    (* Is there a repetition? If so, replace later line by earlier one. *)
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   396
    else case take_prefix (not o is_same_inference (role, t)) lines of
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   397
      (* FIXME: Doesn't this code risk conflating proofs involving different
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   398
         types? *)
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   399
      (_, []) => line :: lines
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   400
    | (pre, Inference_Step (name', _, _, _, _) :: post) =>
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   401
      line :: pre @ map (replace_dependencies_in_line (name', [name])) post
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   402
    | _ => raise Fail "unexpected inference"
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   403
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   404
val waldmeister_conjecture_num = "1.0.0.0"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   405
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   406
val repair_waldmeister_endgame =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   407
  let
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   408
    fun do_tail (Inference_Step (name, _, t, rule, deps)) =
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   409
        Inference_Step (name, Negated_Conjecture, s_not t, rule, deps)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   410
      | do_tail line = line
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   411
    fun do_body [] = []
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   412
      | do_body ((line as Inference_Step ((num, _), _, _, _, _)) :: lines) =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   413
        if num = waldmeister_conjecture_num then map do_tail (line :: lines)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   414
        else line :: do_body lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   415
      | do_body (line :: lines) = line :: do_body lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   416
  in do_body end
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   417
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   418
(* Recursively delete empty lines (type information) from the proof. *)
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   419
fun add_nontrivial_line (line as Inference_Step (name, _, t, _, [])) lines =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   420
    if is_only_type_information t then delete_dependency name lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   421
    else line :: lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   422
  | add_nontrivial_line line lines = line :: lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   423
and delete_dependency name lines =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   424
  fold_rev add_nontrivial_line
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   425
           (map (replace_dependencies_in_line (name, [])) lines) []
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   426
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   427
(* ATPs sometimes reuse free variable names in the strangest ways. Removing
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   428
   offending lines often does the trick. *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   429
fun is_bad_free frees (Free x) = not (member (op =) frees x)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   430
  | is_bad_free _ _ = false
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   431
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   432
val e_skolemize_rule = "skolemize"
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   433
val vampire_skolemisation_rule = "skolemisation"
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   434
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   435
val is_skolemize_rule =
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   436
  member (op =) [e_skolemize_rule, vampire_skolemisation_rule]
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   437
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   438
fun add_desired_line _ _ (line as Definition_Step (name, _, _)) (j, lines) =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   439
    (j, line :: map (replace_dependencies_in_line (name, [])) lines)
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   440
  | add_desired_line fact_names frees
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   441
        (Inference_Step (name as (_, ss), role, t, rule, deps)) (j, lines) =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   442
    (j + 1,
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   443
     if is_fact fact_names ss orelse
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   444
        is_conjecture ss orelse
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   445
        is_skolemize_rule rule orelse
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   446
        (* the last line must be kept *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   447
        j = 0 orelse
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   448
        (not (is_only_type_information t) andalso
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   449
         null (Term.add_tvars t []) andalso
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   450
         not (exists_subterm (is_bad_free frees) t) andalso
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   451
         length deps >= 2 andalso
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   452
         (* kill next to last line, which usually results in a trivial step *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   453
         j <> 1) then
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   454
       Inference_Step (name, role, t, rule, deps) :: lines  (* keep line *)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   455
     else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   456
       map (replace_dependencies_in_line (name, deps)) lines)  (* drop line *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   457
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   458
val indent_size = 2
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   459
val no_label = ("", ~1)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   460
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   461
fun string_for_proof ctxt type_enc lam_trans i n =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   462
  let
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   463
    fun of_indent ind = replicate_string (ind * indent_size) " "
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   464
    fun of_free (s, T) =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   465
      maybe_quote s ^ " :: " ^
50049
dd6a4655cd72 avoid messing too much with output of "string_of_term", so that it doesn't break the yxml encoding for jEdit
blanchet
parents: 50048
diff changeset
   466
      maybe_quote (simplify_spaces (with_vanilla_print_mode
50048
fb024bbf4b65 centralized term printing code
blanchet
parents: 50020
diff changeset
   467
        (Syntax.string_of_typ ctxt) T))
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   468
    val of_frees = space_implode " and " o map of_free
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   469
    fun of_label l = if l = no_label then "" else string_for_label l ^ ": "
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   470
    fun of_have qs =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   471
      (if member (op =) qs Ultimately then "ultimately " else "") ^
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   472
      (if member (op =) qs Then then
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   473
         if member (op =) qs Show then "thus" else "hence"
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   474
       else
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   475
         if member (op =) qs Show then "show" else "have")
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   476
    fun of_obtain qs xs =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   477
      (if member (op =) qs Then then "then " else "") ^ "obtain " ^
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   478
      of_frees xs ^ " where"
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   479
    val of_term =
50048
fb024bbf4b65 centralized term printing code
blanchet
parents: 50020
diff changeset
   480
      annotate_types ctxt
fb024bbf4b65 centralized term printing code
blanchet
parents: 50020
diff changeset
   481
      #> with_vanilla_print_mode (Syntax.string_of_term ctxt)
50049
dd6a4655cd72 avoid messing too much with output of "string_of_term", so that it doesn't break the yxml encoding for jEdit
blanchet
parents: 50048
diff changeset
   482
      #> simplify_spaces
50048
fb024bbf4b65 centralized term printing code
blanchet
parents: 50020
diff changeset
   483
      #> maybe_quote
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   484
    val reconstr = Metis (type_enc, lam_trans)
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   485
    fun of_metis ind options (ls, ss) =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   486
      "\n" ^ of_indent (ind + 1) ^ options ^
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   487
      reconstructor_command reconstr 1 1 [] 0
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   488
          (ls |> sort_distinct (prod_ord string_ord int_ord),
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   489
           ss |> sort_distinct string_ord)
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   490
    and of_step ind (Fix xs) = of_indent ind ^ "fix " ^ of_frees xs ^ "\n"
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   491
      | of_step ind (Let (t1, t2)) =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   492
        of_indent ind ^ "let " ^ of_term t1 ^ " = " ^ of_term t2 ^ "\n"
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   493
      | of_step ind (Assume (l, t)) =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   494
        of_indent ind ^ "assume " ^ of_label l ^ of_term t ^ "\n"
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   495
      | of_step ind (Obtain (qs, xs, l, t, By_Metis facts)) =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   496
        of_indent ind ^ of_obtain qs xs ^ " " ^
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   497
        of_label l ^ of_term t ^
50705
0e943b33d907 use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
blanchet
parents: 50679
diff changeset
   498
        (* The new skolemizer puts the arguments in the same order as the ATPs
0e943b33d907 use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
blanchet
parents: 50679
diff changeset
   499
           (E and Vampire -- but see also "atp_proof_reconstruct.ML" regarding
0e943b33d907 use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
blanchet
parents: 50679
diff changeset
   500
           Vampire). *)
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   501
        of_metis ind "using [[metis_new_skolem]] " facts ^ "\n"
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   502
      | of_step ind (Prove (qs, l, t, By_Metis facts)) =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   503
        of_prove ind qs l t facts
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   504
      | of_step ind (Prove (qs, l, t, Case_Split proofs)) =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   505
        implode (map (prefix (of_indent ind ^ "moreover\n") o of_block ind)
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   506
                     proofs) ^
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   507
        of_prove ind qs l t ([], [])
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   508
      | of_step ind (Prove (qs, l, t, Subblock proof)) =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   509
        of_block ind proof ^ of_prove ind qs l t ([], [])
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   510
    and of_steps prefix suffix ind steps =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   511
      let val s = implode (map (of_step ind) steps) in
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   512
        replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   513
        String.extract (s, ind * indent_size,
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   514
                        SOME (size s - ind * indent_size - 1)) ^
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   515
        suffix ^ "\n"
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   516
      end
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   517
    and of_block ind proof = of_steps "{ " " }" (ind + 1) proof
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   518
    and of_prove ind qs l t facts =
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   519
      of_indent ind ^ of_have qs ^ " " ^ of_label l ^ of_term t ^
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   520
      of_metis ind "" facts ^ "\n"
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   521
    (* One-step proofs are pointless; better use the Metis one-liner
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   522
       directly. *)
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   523
    and of_proof [Prove (_, _, _, By_Metis _)] = ""
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   524
      | of_proof proof =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   525
        (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   526
        of_indent 0 ^ "proof -\n" ^ of_steps "" "" 1 proof ^ of_indent 0 ^
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   527
        (if n <> 1 then "next" else "qed")
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   528
  in of_proof end
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   529
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   530
fun used_labels_of_step (Obtain (_, _, _, _, By_Metis (ls, _))) = ls
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   531
  | used_labels_of_step (Prove (_, _, _, by)) =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   532
    (case by of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   533
       By_Metis (ls, _) => ls
51147
020a6812a204 removed dead weight from data structure
blanchet
parents: 51146
diff changeset
   534
     | Case_Split proofs => fold (union (op =) o used_labels_of) proofs []
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   535
     | Subblock proof => used_labels_of proof)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   536
  | used_labels_of_step _ = []
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   537
and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   538
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   539
fun kill_useless_labels_in_proof proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   540
  let
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   541
    val used_ls = used_labels_of proof
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   542
    fun do_label l = if member (op =) used_ls l then l else no_label
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   543
    fun do_step (Assume (l, t)) = Assume (do_label l, t)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   544
      | do_step (Obtain (qs, xs, l, t, by)) = Obtain (qs, xs, do_label l, t, by)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   545
      | do_step (Prove (qs, l, t, by)) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   546
        Prove (qs, do_label l, t,
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   547
               case by of
51147
020a6812a204 removed dead weight from data structure
blanchet
parents: 51146
diff changeset
   548
                 Case_Split proofs => Case_Split (map do_proof proofs)
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   549
               | Subblock proof => Subblock (do_proof proof)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   550
               | _ => by)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   551
      | do_step step = step
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   552
    and do_proof proof = map do_step proof
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   553
  in do_proof proof end
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   554
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   555
fun prefix_for_depth n = replicate_string (n + 1)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   556
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   557
val relabel_proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   558
  let
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   559
    fun fresh_label depth (old as (l, subst, next_have)) =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   560
      if l = no_label then
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   561
        old
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   562
      else
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   563
        let val l' = (prefix_for_depth depth have_prefix, next_have) in
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   564
          (l', (l, l') :: subst, next_have + 1)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   565
        end
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   566
    fun do_facts subst =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   567
      apfst (maps (the_list o AList.lookup (op =) subst))
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   568
    fun do_byline subst depth nexts by =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   569
      case by of
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   570
        By_Metis facts => By_Metis (do_facts subst facts)
51147
020a6812a204 removed dead weight from data structure
blanchet
parents: 51146
diff changeset
   571
      | Case_Split proofs =>
020a6812a204 removed dead weight from data structure
blanchet
parents: 51146
diff changeset
   572
        Case_Split (map (do_proof subst (depth + 1) (1, 1)) proofs)
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   573
      | Subblock proof => Subblock (do_proof subst depth nexts proof)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   574
    and do_proof _ _ _ [] = []
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   575
      | do_proof subst depth (next_assum, next_have) (Assume (l, t) :: proof) =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   576
        if l = no_label then
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   577
          Assume (l, t) :: do_proof subst depth (next_assum, next_have) proof
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   578
        else
50017
d9c1b11a78d2 avoid name clashes
blanchet
parents: 50016
diff changeset
   579
          let val l' = (prefix_for_depth depth assume_prefix, next_assum) in
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   580
            Assume (l', t) ::
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   581
            do_proof ((l, l') :: subst) depth (next_assum + 1, next_have) proof
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   582
          end
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   583
      | do_proof subst depth (nexts as (next_assum, next_have))
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   584
            (Obtain (qs, xs, l, t, by) :: proof) =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   585
        let
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   586
          val (l, subst, next_have) = (l, subst, next_have) |> fresh_label depth
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   587
          val by = by |> do_byline subst depth nexts
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   588
        in
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   589
          Obtain (qs, xs, l, t, by) ::
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   590
          do_proof subst depth (next_assum, next_have) proof
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   591
        end
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   592
      | do_proof subst depth (nexts as (next_assum, next_have))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   593
            (Prove (qs, l, t, by) :: proof) =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   594
        let
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   595
          val (l, subst, next_have) = (l, subst, next_have) |> fresh_label depth
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   596
          val by = by |> do_byline subst depth nexts
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   597
        in
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   598
          Prove (qs, l, t, by) ::
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   599
          do_proof subst depth (next_assum, next_have) proof
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   600
        end
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   601
      | do_proof subst depth nexts (step :: proof) =
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   602
        step :: do_proof subst depth nexts proof
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   603
  in do_proof [] 0 (1, 1) end
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   604
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   605
val chain_direct_proof =
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   606
  let
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   607
    fun label_of (Assume (l, _)) = SOME l
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   608
      | label_of (Obtain (_, _, l, _, _)) = SOME l
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   609
      | label_of (Prove (_, l, _, _)) = SOME l
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   610
      | label_of _ = NONE
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   611
    fun chain_step (SOME l0)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   612
                   (step as Obtain (qs, xs, l, t, By_Metis (lfs, gfs))) =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   613
        if member (op =) lfs l0 then
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   614
          Obtain (Then :: qs, xs, l, t, By_Metis (lfs |> remove (op =) l0, gfs))
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   615
        else
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   616
          step
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   617
      | chain_step (SOME l0)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   618
                   (step as Prove (qs, l, t, By_Metis (lfs, gfs))) =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   619
        if member (op =) lfs l0 then
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   620
          Prove (Then :: qs, l, t, By_Metis (lfs |> remove (op =) l0, gfs))
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   621
        else
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   622
          step
51147
020a6812a204 removed dead weight from data structure
blanchet
parents: 51146
diff changeset
   623
      | chain_step _ (Prove (qs, l, t, Case_Split proofs)) =
020a6812a204 removed dead weight from data structure
blanchet
parents: 51146
diff changeset
   624
        Prove (qs, l, t, Case_Split (map (chain_proof NONE) proofs))
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   625
      | chain_step _ (Prove (qs, l, t, Subblock proof)) =
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   626
        Prove (qs, l, t, Subblock (chain_proof NONE proof))
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   627
      | chain_step _ step = step
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   628
    and chain_proof _ [] = []
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   629
      | chain_proof (prev as SOME _) (i :: is) =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   630
        chain_step prev i :: chain_proof (label_of i) is
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   631
      | chain_proof _ (i :: is) = i :: chain_proof (label_of i) is
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   632
  in chain_proof NONE end
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   633
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   634
type isar_params =
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50450
diff changeset
   635
  bool * bool * Time.time option * real * string Symtab.table
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   636
  * (string * stature) list vector * int Symtab.table * string proof * thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   637
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   638
fun isar_proof_text ctxt isar_proofs
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51129
diff changeset
   639
    (debug, verbose, preplay_timeout, isar_compress, pool, fact_names, sym_tab,
50020
6b9611abcd4c renamed Sledgehammer option
blanchet
parents: 50019
diff changeset
   640
     atp_proof, goal)
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   641
    (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   642
  let
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   643
    val (params, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   644
    val frees = fold Term.add_frees (concl_t :: hyp_ts) []
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   645
    val one_line_proof = one_line_proof_text 0 one_line_params
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   646
    val type_enc =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   647
      if is_typed_helper_used_in_atp_proof atp_proof then full_typesN
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   648
      else partial_typesN
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   649
    val lam_trans = lam_trans_from_atp_proof atp_proof metis_default_lam_trans
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50450
diff changeset
   650
    val preplay = preplay_timeout <> SOME Time.zeroTime
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   651
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   652
    fun isar_proof_of () =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   653
      let
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   654
        val atp_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   655
          atp_proof
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   656
          |> clean_up_atp_proof_dependencies
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   657
          |> nasty_atp_proof pool
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   658
          |> map_term_names_in_atp_proof repair_name
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   659
          |> decode_lines ctxt sym_tab
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   660
          |> repair_waldmeister_endgame
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   661
          |> rpair [] |-> fold_rev (add_line fact_names)
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   662
          |> rpair [] |-> fold_rev add_nontrivial_line
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   663
          |> rpair (0, [])
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   664
          |-> fold_rev (add_desired_line fact_names frees)
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   665
          |> snd
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   666
        val conj_name = conjecture_prefix ^ string_of_int (length hyp_ts)
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   667
        val conjs =
50010
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   668
          atp_proof |> map_filter
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   669
            (fn Inference_Step (name as (_, ss), _, _, _, []) =>
50010
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   670
                if member (op =) ss conj_name then SOME name else NONE
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   671
              | _ => NONE)
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   672
        val assms =
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   673
          atp_proof |> map_filter
50013
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   674
            (fn Inference_Step (name as (_, ss), _, _, _, []) =>
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   675
                (case resolve_conjecture ss of
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   676
                   [j] =>
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   677
                   if j = length hyp_ts then NONE
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   678
                   else SOME (Assume (raw_label_for_name name, nth hyp_ts j))
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   679
                 | _ => NONE)
50010
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   680
              | _ => NONE)
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   681
        fun dep_of_step (Definition_Step _) = NONE
50012
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   682
          | dep_of_step (Inference_Step (name, _, _, _, from)) =
01cb92151a53 track formula roles in proofs and use that to determine whether the conjecture should be negated or not
blanchet
parents: 50010
diff changeset
   683
            SOME (from, name)
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   684
        val refute_graph =
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   685
          atp_proof |> map_filter dep_of_step |> make_refute_graph
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   686
        val axioms = axioms_of_refute_graph refute_graph conjs
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   687
        val tainted = tainted_atoms_of_refute_graph refute_graph conjs
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   688
        val is_clause_tainted = exists (member (op =) tainted)
51031
63d71b247323 more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
blanchet
parents: 51026
diff changeset
   689
        val bot = atp_proof |> List.last |> dep_of_step |> the |> snd
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   690
        val steps =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   691
          Symtab.empty
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   692
          |> fold (fn Definition_Step _ => I (* FIXME *)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   693
                    | Inference_Step (name as (s, ss), role, t, rule, _) =>
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   694
                      Symtab.update_new (s, (rule,
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   695
                        t |> (if is_clause_tainted [name] then
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   696
                                s_maybe_not role
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   697
                                #> fold exists_of (map Var (Term.add_vars t []))
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   698
                              else
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   699
                                I))))
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   700
                  atp_proof
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   701
        fun is_clause_skolemize_rule [(s, _)] =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   702
            Option.map (is_skolemize_rule o fst) (Symtab.lookup steps s) =
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   703
            SOME true
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   704
          | is_clause_skolemize_rule _ = false
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   705
        (* The assumptions and conjecture are "prop"s; the other formulas are
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   706
           "bool"s. *)
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   707
        fun prop_of_clause [name as (s, ss)] =
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   708
            (case resolve_conjecture ss of
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   709
               [j] => if j = length hyp_ts then concl_t else nth hyp_ts j
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   710
             | _ => the_default ("", @{term False}) (Symtab.lookup steps s)
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   711
                    |> snd |> HOLogic.mk_Trueprop |> close_form)
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   712
          | prop_of_clause names =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   713
            let
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   714
              val lits = map snd (map_filter (Symtab.lookup steps o fst) names)
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   715
            in
50018
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   716
              case List.partition (can HOLogic.dest_not) lits of
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   717
                (negs as _ :: _, pos as _ :: _) =>
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   718
                HOLogic.mk_imp
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   719
                  (Library.foldr1 s_conj (map HOLogic.dest_not negs),
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   720
                   Library.foldr1 s_disj pos)
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   721
              | _ => fold (curry s_disj) lits @{term False}
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   722
            end
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   723
            |> HOLogic.mk_Trueprop |> close_form
51164
eba05aaa8612 repaired collateral damage from 4f0147ed8bcb
blanchet
parents: 51158
diff changeset
   724
        fun isar_proof_of_direct_proof outer accum [] =
51165
blanchet
parents: 51164
diff changeset
   725
            rev accum |> outer ? (append assms
blanchet
parents: 51164
diff changeset
   726
                                  #> not (null params) ? cons (Fix params))
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   727
          | isar_proof_of_direct_proof outer accum (inf :: infs) =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   728
            let
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   729
              fun maybe_show outer c =
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   730
                (outer andalso length c = 1 andalso subset (op =) (c, conjs))
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   731
                ? cons Show
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   732
              fun do_rest step =
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   733
                isar_proof_of_direct_proof outer (step :: accum) infs
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   734
              val is_fixed = Variable.is_declared ctxt orf can Name.dest_skolem
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   735
              fun skolems_of t =
51158
f432363eebf4 annotate obtains with types
blanchet
parents: 51156
diff changeset
   736
                Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   737
            in
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   738
              case inf of
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   739
                Have (gamma, c) =>
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   740
                let
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   741
                  val l = label_of_clause c
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   742
                  val t = prop_of_clause c
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   743
                  val by =
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   744
                    By_Metis (fold (add_fact_from_dependencies fact_names) gamma
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   745
                                   ([], []))
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   746
                  fun prove outer = Prove (maybe_show outer c [], l, t, by)
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   747
                in
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   748
                  if is_clause_tainted c then
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   749
                    case gamma of
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   750
                      [g] =>
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   751
                      if is_clause_skolemize_rule g andalso
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   752
                         is_clause_tainted g then
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   753
                        let
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   754
                          val proof =
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   755
                            Fix (skolems_of (prop_of_clause g)) :: rev accum
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   756
                        in
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   757
                          isar_proof_of_direct_proof outer
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   758
                              [Prove (maybe_show outer c [Then],
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   759
                                      label_of_clause c, prop_of_clause c,
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   760
                                      Subblock proof)] []
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   761
                        end
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   762
                      else
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   763
                        do_rest (prove outer)
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   764
                    | _ => do_rest (prove outer)
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   765
                  else
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   766
                    if is_clause_skolemize_rule c then
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   767
                      do_rest (Obtain ([], skolems_of t, l, t, by))
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   768
                    else
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   769
                      do_rest (prove outer)
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   770
                end
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   771
              | Cases cases =>
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   772
                let
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   773
                  fun do_case (c, infs) =
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   774
                    Assume (label_of_clause c, prop_of_clause c) ::
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   775
                    isar_proof_of_direct_proof false [] infs
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   776
                  val c = succedent_of_cases cases
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   777
                in
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   778
                  do_rest (Prove (maybe_show outer c [Ultimately],
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   779
                                  label_of_clause c, prop_of_clause c,
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   780
                                  Case_Split (map do_case cases)))
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   781
                end
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   782
            end
51165
blanchet
parents: 51164
diff changeset
   783
        val cleanup_labels_in_proof =
blanchet
parents: 51164
diff changeset
   784
          chain_direct_proof
blanchet
parents: 51164
diff changeset
   785
          #> kill_useless_labels_in_proof
blanchet
parents: 51164
diff changeset
   786
          #> relabel_proof
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50905
diff changeset
   787
        val (isar_proof, (preplay_fail, preplay_time)) =
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   788
          refute_graph
51031
63d71b247323 more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
blanchet
parents: 51026
diff changeset
   789
          |> redirect_graph axioms tainted bot
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   790
          |> isar_proof_of_direct_proof true []
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51129
diff changeset
   791
          |> (if not preplay andalso isar_compress <= 1.0 then
50677
f5c217474eca use rpair to avoid swap
smolkas
parents: 50676
diff changeset
   792
                rpair (false, (true, seconds 0.0))
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50450
diff changeset
   793
              else
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51129
diff changeset
   794
                compress_proof debug ctxt type_enc lam_trans preplay
50679
e409d9f8ec75 removed whitespace
smolkas
parents: 50677
diff changeset
   795
                  preplay_timeout
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51129
diff changeset
   796
                  (if isar_proofs then isar_compress else 1000.0))
51165
blanchet
parents: 51164
diff changeset
   797
          |>> cleanup_labels_in_proof
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   798
        val isar_text =
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   799
          string_for_proof ctxt type_enc lam_trans subgoal subgoal_count
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   800
                           isar_proof
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   801
      in
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   802
        case isar_text of
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   803
          "" =>
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   804
          if isar_proofs then
50671
blanchet
parents: 50670
diff changeset
   805
            "\nNo structured proof available (proof too simple)."
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   806
          else
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   807
            ""
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   808
        | _ =>
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   809
          let
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   810
            val msg =
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   811
              (if preplay then
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50905
diff changeset
   812
                [(if preplay_fail then "may fail, " else "") ^
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50905
diff changeset
   813
                   Sledgehammer_Preplay.string_of_preplay_time preplay_time]
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   814
               else
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   815
                 []) @
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   816
              (if verbose then
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   817
                 let val num_steps = metis_steps_total isar_proof in
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   818
                   [string_of_int num_steps ^ " step" ^ plural_s num_steps]
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   819
                 end
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   820
               else
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   821
                 [])
50277
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   822
          in
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   823
            "\n\nStructured proof "
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   824
              ^ (commas msg |> not (null msg) ? enclose "(" ")")
50450
358b6020f8b6 generalized notion of active area, where sendback is just one application;
wenzelm
parents: 50410
diff changeset
   825
              ^ ":\n" ^ Active.sendback_markup isar_text
50277
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   826
          end
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   827
      end
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   828
    val isar_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   829
      if debug then
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   830
        isar_proof_of ()
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   831
      else case try isar_proof_of () of
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   832
        SOME s => s
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   833
      | NONE => if isar_proofs then
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   834
                  "\nWarning: The Isar proof construction failed."
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   835
                else
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   836
                  ""
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   837
  in one_line_proof ^ isar_proof end
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   838
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   839
fun proof_text ctxt isar_proofs isar_params num_chained
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   840
               (one_line_params as (preplay, _, _, _, _, _)) =
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   841
  (if case preplay of Failed_to_Play _ => true | _ => isar_proofs then
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   842
     isar_proof_text ctxt isar_proofs isar_params
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   843
   else
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   844
     one_line_proof_text num_chained) one_line_params
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   845
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   846
end;