src/HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML
author blanchet
Sun, 02 Feb 2014 20:53:51 +0100
changeset 55260 ada3ae6458d4
parent 55257 abfd7b90bba2
child 55278 73372494fd80
permissions -rw-r--r--
more data structure rationalization
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML
50263
0b430064296a added comments to new source files
smolkas
parents: 50260
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
0b430064296a added comments to new source files
smolkas
parents: 50260
diff changeset
     3
    Author:     Steffen Juilf Smolka, TU Muenchen
0b430064296a added comments to new source files
smolkas
parents: 50260
diff changeset
     4
0b430064296a added comments to new source files
smolkas
parents: 50260
diff changeset
     5
Basic data structures for representing and basic methods
0b430064296a added comments to new source files
smolkas
parents: 50260
diff changeset
     6
for dealing with Isar proof texts.
0b430064296a added comments to new source files
smolkas
parents: 50260
diff changeset
     7
*)
0b430064296a added comments to new source files
smolkas
parents: 50260
diff changeset
     8
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
     9
signature SLEDGEHAMMER_ISAR_PROOF =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    10
sig
51239
67cc209493b2 eliminated hard tabs;
wenzelm
parents: 51179
diff changeset
    11
  type label = string * int
54816
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54767
diff changeset
    12
  type facts = label list * string list (* local and global facts *)
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
    13
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
    14
  datatype isar_qualifier = Show | Then
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
    15
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
    16
  datatype proof_method =
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    17
    Metis_Method of string option * string option |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    18
    Simp_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    19
    Simp_Size_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    20
    Auto_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    21
    Fastforce_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    22
    Force_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    23
    Arith_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    24
    Blast_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    25
    Meson_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    26
    Algebra_Method
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
    27
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    28
  datatype isar_proof =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54534
diff changeset
    29
    Proof of (string * typ) list * (label * term) list * isar_step list
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    30
  and isar_step =
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
    31
    Let of term * term |
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54534
diff changeset
    32
    Prove of isar_qualifier list * (string * typ) list * label * term * isar_proof list *
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
    33
      (facts * proof_method list)
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    34
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    35
  val no_label : label
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    36
  val no_facts : facts
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
    37
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    38
  val label_ord : label * label -> order
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51239
diff changeset
    39
  val string_of_label : label -> string
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    40
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
    41
  val string_of_proof_method : proof_method -> string
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
    42
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55257
diff changeset
    43
  val steps_of_isar_proof : isar_proof -> isar_step list
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    44
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
    45
  val label_of_isar_step : isar_step -> label option
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
    46
  val byline_of_isar_step : isar_step -> (facts * proof_method list) option
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    47
54765
blanchet
parents: 54764
diff changeset
    48
  val fold_isar_steps : (isar_step -> 'a -> 'a) -> isar_step list -> 'a -> 'a
55212
blanchet
parents: 55211
diff changeset
    49
  val map_isar_steps : (isar_step -> isar_step) -> isar_proof -> isar_proof
blanchet
parents: 55211
diff changeset
    50
  val add_isar_steps : isar_step list -> int -> int
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    51
55212
blanchet
parents: 55211
diff changeset
    52
  structure Canonical_Label_Tab : TABLE
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    53
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    54
  val canonical_label_ord : (label * label) -> order
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
    55
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
    56
  val chain_isar_proof : isar_proof -> isar_proof
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
    57
  val kill_useless_labels_in_isar_proof : isar_proof -> isar_proof
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
    58
  val relabel_isar_proof_canonically : isar_proof -> isar_proof
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
    59
  val relabel_isar_proof_finally : isar_proof -> isar_proof
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    60
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    61
  val string_of_isar_proof : Proof.context -> int -> int ->
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
    62
    (label -> proof_method list -> string) -> isar_proof -> string
54504
blanchet
parents: 54503
diff changeset
    63
end;
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    64
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    65
structure Sledgehammer_Isar_Proof : SLEDGEHAMMER_ISAR_PROOF =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    66
struct
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    67
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    68
open ATP_Util
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    69
open ATP_Proof
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    70
open ATP_Problem_Generate
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    71
open ATP_Proof_Reconstruct
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    72
open Sledgehammer_Util
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    73
open Sledgehammer_Reconstructor
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    74
open Sledgehammer_Isar_Annotate
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
    75
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    76
type label = string * int
54534
blanchet
parents: 54504
diff changeset
    77
type facts = label list * string list (* local and global facts *)
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    78
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
    79
datatype isar_qualifier = Show | Then
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    80
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
    81
datatype proof_method =
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    82
  Metis_Method of string option * string option |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    83
  Simp_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    84
  Simp_Size_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    85
  Auto_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    86
  Fastforce_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    87
  Force_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    88
  Arith_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    89
  Blast_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    90
  Meson_Method |
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
    91
  Algebra_Method
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
    92
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    93
datatype isar_proof =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54534
diff changeset
    94
  Proof of (string * typ) list * (label * term) list * isar_step list
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    95
and isar_step =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    96
  Let of term * term |
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54534
diff changeset
    97
  Prove of isar_qualifier list * (string * typ) list * label * term * isar_proof list *
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
    98
    (facts * proof_method list)
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    99
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   100
val no_label = ("", ~1)
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   101
val no_facts = ([],[])
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   102
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   103
val label_ord = pairself swap #> prod_ord int_ord fast_string_ord
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   104
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51239
diff changeset
   105
fun string_of_label (s, num) = s ^ string_of_int num
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   106
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   107
fun string_of_proof_method meth =
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   108
  (case meth of
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   109
    Metis_Method (NONE, NONE) => "metis"
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   110
  | Metis_Method (type_enc_opt, lam_trans_opt) =>
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   111
    "metis (" ^ commas (map_filter I [type_enc_opt, lam_trans_opt]) ^ ")"
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   112
  | Simp_Method => "simp"
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   113
  | Simp_Size_Method => "simp add: size_ne_size_imp_ne"
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   114
  | Auto_Method => "auto"
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   115
  | Fastforce_Method => "fastforce"
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   116
  | Force_Method => "force"
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   117
  | Arith_Method => "arith"
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   118
  | Blast_Method => "blast"
55219
6fe9fd75f9d7 added 'algebra' to the mix
blanchet
parents: 55218
diff changeset
   119
  | Meson_Method => "meson"
6fe9fd75f9d7 added 'algebra' to the mix
blanchet
parents: 55218
diff changeset
   120
  | Algebra_Method => "algebra")
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   121
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55257
diff changeset
   122
fun steps_of_isar_proof (Proof (_, _, steps)) = steps
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   123
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   124
fun label_of_isar_step (Prove (_, _, l, _, _, _)) = SOME l
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   125
  | label_of_isar_step _ = NONE
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   126
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   127
fun subproofs_of_isar_step (Prove (_, _, _, _, subproofs, _)) = SOME subproofs
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   128
  | subproofs_of_isar_step _ = NONE
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   129
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   130
fun byline_of_isar_step (Prove (_, _, _, _, _, byline)) = SOME byline
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   131
  | byline_of_isar_step _ = NONE
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   132
55212
blanchet
parents: 55211
diff changeset
   133
fun fold_isar_step f step =
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55257
diff changeset
   134
  fold (steps_of_isar_proof #> fold_isar_steps f) (these (subproofs_of_isar_step step)) #> f step
55212
blanchet
parents: 55211
diff changeset
   135
and fold_isar_steps f = fold (fold_isar_step f)
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   136
54765
blanchet
parents: 54764
diff changeset
   137
fun map_isar_steps f =
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   138
  let
55212
blanchet
parents: 55211
diff changeset
   139
    fun map_proof (Proof (fix, assms, steps)) = Proof (fix, assms, map map_step steps)
blanchet
parents: 55211
diff changeset
   140
    and map_step (step as Let _) = f step
blanchet
parents: 55211
diff changeset
   141
      | map_step (Prove (qs, xs, l, t, subproofs, by)) =
blanchet
parents: 55211
diff changeset
   142
        f (Prove (qs, xs, l, t, map map_proof subproofs, by))
blanchet
parents: 55211
diff changeset
   143
  in map_proof end
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   144
55212
blanchet
parents: 55211
diff changeset
   145
val add_isar_steps = fold_isar_steps (fn Prove _ => Integer.add 1 | _ => I)
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   146
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   147
(* canonical proof labels: 1, 2, 3, ... in post traversal order *)
52557
92ed2926596d made SML/NJ happy
smolkas
parents: 52556
diff changeset
   148
fun canonical_label_ord (((_, i1), (_, i2)) : label * label) = int_ord (i1, i2)
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   149
55212
blanchet
parents: 55211
diff changeset
   150
structure Canonical_Label_Tab = Table(
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   151
  type key = label
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   152
  val ord = canonical_label_ord)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   153
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   154
fun chain_qs_lfs NONE lfs = ([], lfs)
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   155
  | chain_qs_lfs (SOME l0) lfs =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   156
    if member (op =) lfs l0 then ([Then], lfs |> remove (op =) l0) else ([], lfs)
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   157
fun chain_isar_step lbl (Prove (qs, xs, l, t, subproofs, ((lfs, gfs), meths))) =
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   158
    let val (qs', lfs) = chain_qs_lfs lbl lfs in
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   159
      Prove (qs' @ qs, xs, l, t, map chain_isar_proof subproofs, ((lfs, gfs), meths))
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   160
    end
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   161
  | chain_isar_step _ step = step
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   162
and chain_isar_steps _ [] = []
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   163
  | chain_isar_steps (prev as SOME _) (i :: is) =
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   164
    chain_isar_step prev i :: chain_isar_steps (label_of_isar_step i) is
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   165
  | chain_isar_steps _ (i :: is) = i :: chain_isar_steps (label_of_isar_step i) is
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   166
and chain_isar_proof (Proof (fix, assms, steps)) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   167
  Proof (fix, assms, chain_isar_steps (try (List.last #> fst) assms) steps)
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   168
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   169
fun kill_useless_labels_in_isar_proof proof =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   170
  let
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   171
    val used_ls =
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   172
      fold_isar_steps (byline_of_isar_step #> (fn SOME ((ls, _), _) => union (op =) ls | _ => I))
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55257
diff changeset
   173
        (steps_of_isar_proof proof) []
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   174
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   175
    fun kill_label l = if member (op =) used_ls l then l else no_label
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   176
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   177
    fun kill_step (Prove (qs, xs, l, t, subproofs, by)) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   178
        Prove (qs, xs, kill_label l, t, map kill_proof subproofs, by)
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   179
      | kill_step step = step
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   180
    and kill_proof (Proof (fix, assms, steps)) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   181
      Proof (fix, map (apfst kill_label) assms, map kill_step steps)
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   182
  in
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   183
    kill_proof proof
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   184
  end
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   185
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   186
fun relabel_isar_proof_canonically proof =
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   187
  let
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   188
    fun next_label l (next, subst) =
54534
blanchet
parents: 54504
diff changeset
   189
      let val l' = ("", next) in (l', (next + 1, (l, l') :: subst)) end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   190
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54534
diff changeset
   191
    fun do_byline by (_, subst) = apfst (apfst (map (AList.lookup (op =) subst #> the))) by
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   192
    handle Option.Option =>
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   193
      raise Fail "Sledgehammer_Isar_Proof: relabel_isar_proof_canonically"
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   194
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   195
    fun do_assm (l, t) state =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   196
      let
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   197
        val (l, state) = next_label l state
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   198
      in ((l, t), state) end
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   199
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54534
diff changeset
   200
    fun do_proof (Proof (fix, assms, steps)) state =
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   201
      let
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   202
        val (assms, state) = fold_map do_assm assms state
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   203
        val (steps, state) = fold_map do_step steps state
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   204
      in
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54534
diff changeset
   205
        (Proof (fix, assms, steps), state)
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   206
      end
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   207
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   208
    and do_step (step as Let _) state = (step, state)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   209
      | do_step (Prove (qs, fix, l, t, subproofs, by)) state=
54534
blanchet
parents: 54504
diff changeset
   210
        let
blanchet
parents: 54504
diff changeset
   211
          val by = do_byline by state
blanchet
parents: 54504
diff changeset
   212
          val (subproofs, state) = fold_map do_proof subproofs state
blanchet
parents: 54504
diff changeset
   213
          val (l, state) = next_label l state
blanchet
parents: 54504
diff changeset
   214
        in
blanchet
parents: 54504
diff changeset
   215
          (Prove (qs, fix, l, t, subproofs, by), state)
blanchet
parents: 54504
diff changeset
   216
        end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   217
  in
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   218
    fst (do_proof proof (0, []))
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   219
  end
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   220
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   221
val assume_prefix = "a"
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   222
val have_prefix = "f"
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   223
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   224
val relabel_isar_proof_finally =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   225
  let
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   226
    fun fresh_label depth prefix (accum as (l, subst, next)) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   227
      if l = no_label then
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   228
        accum
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   229
      else
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   230
        let val l' = (replicate_string (depth + 1) prefix, next) in
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   231
          (l', (l, l') :: subst, next + 1)
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   232
        end
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   233
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   234
    fun relabel_facts subst = apfst (maps (the_list o AList.lookup (op =) subst))
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   235
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   236
    fun relabel_assm depth (l, t) (subst, next) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   237
      let val (l, subst, next) = (l, subst, next) |> fresh_label depth assume_prefix in
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   238
        ((l, t), (subst, next))
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   239
      end
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   240
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   241
    fun relabel_assms subst depth assms = fold_map (relabel_assm depth) assms (subst, 1) ||> fst
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   242
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   243
    fun relabel_steps _ _ _ [] = []
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   244
      | relabel_steps subst depth next (Prove (qs, xs, l, t, sub, by) :: steps) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   245
        let
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   246
          val (l, subst, next) = (l, subst, next) |> fresh_label depth have_prefix
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   247
          val sub = relabel_proofs subst depth sub
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   248
          val by = apfst (relabel_facts subst) by
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   249
        in
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   250
          Prove (qs, xs, l, t, sub, by) :: relabel_steps subst depth next steps
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   251
        end
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   252
      | relabel_steps subst depth next (step :: steps) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   253
        step :: relabel_steps subst depth next steps
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   254
    and relabel_proof subst depth (Proof (fix, assms, steps)) =
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   255
      let val (assms, subst) = relabel_assms subst depth assms in
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   256
        Proof (fix, assms, relabel_steps subst depth 1 steps)
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   257
      end
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   258
    and relabel_proofs subst depth = map (relabel_proof subst (depth + 1))
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   259
  in
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   260
    relabel_proof [] 0
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   261
  end
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   262
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   263
val indent_size = 2
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   264
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   265
fun string_of_isar_proof ctxt i n comment_of proof =
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   266
  let
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   267
    (* Make sure only type constraints inserted by the type annotation code are printed. *)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   268
    val ctxt =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   269
      ctxt |> Config.put show_markup false
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   270
           |> Config.put Printer.show_type_emphasis false
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   271
           |> Config.put show_types false
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   272
           |> Config.put show_sorts false
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   273
           |> Config.put show_consts false
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   274
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   275
    val register_fixes = map Free #> fold Variable.auto_fixes
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   276
55216
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   277
    fun add_str s' = apfst (suffix s')
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   278
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   279
    fun of_indent ind = replicate_string (ind * indent_size) " "
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   280
    fun of_moreover ind = of_indent ind ^ "moreover\n"
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   281
    fun of_label l = if l = no_label then "" else string_of_label l ^ ": "
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   282
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   283
    fun of_obtain qs nr =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   284
      (if nr > 1 orelse (nr = 1 andalso member (op =) qs Then) then "ultimately "
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   285
       else if nr = 1 orelse member (op =) qs Then then "then "
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   286
       else "") ^ "obtain"
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   287
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   288
    fun of_show_have qs = if member (op =) qs Show then "show" else "have"
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   289
    fun of_thus_hence qs = if member (op =) qs Show then "thus" else "hence"
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   290
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   291
    fun of_have qs nr =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   292
      if nr > 1 orelse (nr = 1 andalso member (op =) qs Then) then "ultimately " ^ of_show_have qs
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   293
      else if nr = 1 orelse member (op =) qs Then then of_thus_hence qs
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   294
      else of_show_have qs
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   295
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   296
    fun add_term term (s, ctxt) =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   297
      (s ^ (term
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   298
            |> singleton (Syntax.uncheck_terms ctxt)
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   299
            |> annotate_types_in_term ctxt
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   300
            |> with_vanilla_print_mode (Syntax.unparse_term ctxt #> Pretty.string_of)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   301
            |> simplify_spaces
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   302
            |> maybe_quote),
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   303
       ctxt |> Variable.auto_fixes term)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   304
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   305
    fun with_facts none _ [] [] = none
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   306
      | with_facts _ some ls ss = some (space_implode " " (map string_of_label ls @ ss))
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   307
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   308
    val using_facts = with_facts "" (enclose "using " " ")
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   309
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   310
    fun maybe_paren s = s |> String.isSubstring " " s ? enclose "(" ")"
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   311
    fun by_facts meth ls ss =
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   312
      "by " ^ with_facts (maybe_paren meth) (enclose ("(" ^ meth ^ " ") ")") ls ss
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   313
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   314
    fun is_direct_method (Metis_Method _) = true
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   315
      | is_direct_method Meson_Method = true
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   316
      | is_direct_method _ = false
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   317
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   318
    (* Local facts are always passed via "using", which affects "meson" and "metis". This is
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   319
       arguably stylistically superior, because it emphasises the structure of the proof. It is also
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   320
       more robust w.r.t. preplay: Preplay is performed before chaining of local facts with "hence"
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   321
       and "thus" is introduced. See also "tac_of_method" in "Sledgehammer_Isar_Preplay". *)
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   322
    fun of_method ls ss meth =
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   323
        let val direct = is_direct_method meth in
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   324
          using_facts ls (if direct then [] else ss) ^
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   325
          by_facts (string_of_proof_method meth) [] (if direct then ss else [])
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   326
        end
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   327
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   328
    fun of_free (s, T) =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   329
      maybe_quote s ^ " :: " ^
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   330
      maybe_quote (simplify_spaces (with_vanilla_print_mode (Syntax.string_of_typ ctxt) T))
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   331
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   332
    fun add_frees xs (s, ctxt) =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   333
      (s ^ space_implode " and " (map of_free xs), ctxt |> register_fixes xs)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   334
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   335
    fun add_fix _ [] = I
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   336
      | add_fix ind xs =
55216
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   337
        add_str (of_indent ind ^ "fix ")
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   338
        #> add_frees xs
55216
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   339
        #> add_str "\n"
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   340
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   341
    fun add_assm ind (l, t) =
55216
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   342
      add_str (of_indent ind ^ "assume " ^ of_label l)
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   343
      #> add_term t
55216
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   344
      #> add_str "\n"
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   345
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   346
    fun add_assms ind assms = fold (add_assm ind) assms
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   347
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   348
    fun of_subproof ind ctxt proof =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   349
      let
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   350
        val ind = ind + 1
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   351
        val s = of_proof ind ctxt proof
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   352
        val prefix = "{ "
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   353
        val suffix = " }"
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   354
      in
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   355
        replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   356
        String.extract (s, ind * indent_size, SOME (size s - ind * indent_size - 1)) ^
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   357
        suffix ^ "\n"
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   358
      end
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   359
    and of_subproofs _ _ _ [] = ""
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   360
      | of_subproofs ind ctxt qs subproofs =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   361
        (if member (op =) qs Then then of_moreover ind else "") ^
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   362
        space_implode (of_moreover ind) (map (of_subproof ind ctxt) subproofs)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   363
    and add_step_pre ind qs subproofs (s, ctxt) =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   364
      (s ^ of_subproofs ind ctxt qs subproofs ^ of_indent ind, ctxt)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   365
    and add_step ind (Let (t1, t2)) =
55216
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   366
        add_str (of_indent ind ^ "let ")
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   367
        #> add_term t1 #> add_str " = " #> add_term t2
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   368
        #> add_str "\n"
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   369
      | add_step ind (Prove (qs, xs, l, t, subproofs, ((ls, ss), meths as meth :: _))) =
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   370
        add_step_pre ind qs subproofs
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   371
        #> (case xs of
55216
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   372
             [] => add_str (of_have qs (length subproofs) ^ " ")
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   373
           | _ =>
77953325d5f1 more concise Isar output
blanchet
parents: 55214
diff changeset
   374
             add_str (of_obtain qs (length subproofs) ^ " ") #> add_frees xs #> add_str " where ")
55217
blanchet
parents: 55216
diff changeset
   375
        #> add_str (of_label l)
blanchet
parents: 55216
diff changeset
   376
        #> add_term t
blanchet
parents: 55216
diff changeset
   377
        #> add_str (" " ^
55218
ed495a5088c6 more informative trace
blanchet
parents: 55217
diff changeset
   378
             of_method (sort_distinct label_ord ls) (sort_distinct string_ord ss) meth ^
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   379
             (case comment_of l meths of
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   380
               "" => ""
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   381
             | comment => " (* " ^ comment ^ " *)") ^ "\n")
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   382
    and add_steps ind = fold (add_step ind)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   383
    and of_proof ind ctxt (Proof (xs, assms, steps)) =
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   384
      ("", ctxt) |> add_fix ind xs |> add_assms ind assms |> add_steps ind steps |> fst
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   385
  in
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   386
    (* One-step Metis proofs are pointless; better use the one-liner directly. *)
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   387
    (case proof of
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   388
      Proof ([], [], []) => "" (* degenerate case: the conjecture is "True" with Z3 *)
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55244
diff changeset
   389
    | Proof ([], [], [Prove (_, [], _, _, [], (_, Metis_Method _ :: _))]) => ""
55211
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   390
    | _ =>
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   391
      (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   392
      of_indent 0 ^ "proof -\n" ^ of_proof 1 ctxt proof ^
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   393
      of_indent 0 ^ (if n <> 1 then "next" else "qed"))
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   394
  end
5d027af93a08 moved ML code around
blanchet
parents: 55202
diff changeset
   395
54504
blanchet
parents: 54503
diff changeset
   396
end;