src/HOL/Tools/Sledgehammer/sledgehammer_proof.ML
author wenzelm
Mon, 05 Aug 2013 15:03:52 +0200
changeset 52861 e93d73b51fd0
parent 52629 d6f2a7c196f7
child 54503 b490e15a5e19
permissions -rw-r--r--
commands with overlay remain visible, to avoid loosing printed output;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50264
a9ec48b98734 renamed sledgehammer_isar_reconstruct to sledgehammer_proof
smolkas
parents: 50263
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_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
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
     9
signature SLEDGEHAMMER_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
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    12
  type facts = label list * string list (* local & 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
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
    16
  datatype fix = Fix of (string * typ) list
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
    17
  datatype assms = Assume of (label * term) list
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
    18
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    19
  datatype isar_proof =
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
    20
    Proof of fix * assms * isar_step list
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
    21
  and isar_step =
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
    22
    Let of term * term |
52590
smolkas
parents: 52557
diff changeset
    23
    (* for |fix|>0, this is an obtain step; step may contain subproofs *)
smolkas
parents: 52557
diff changeset
    24
    Prove of isar_qualifier list * fix * label * term * isar_proof list * byline
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
    25
  and byline =
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    26
    By of facts * proof_method
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    27
  and proof_method =
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    28
    MetisM |
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    29
    SimpM |
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    30
    AutoM |
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    31
    FastforceM |
52629
d6f2a7c196f7 added blast, force
smolkas
parents: 52627
diff changeset
    32
    ForceM |
d6f2a7c196f7 added blast, force
smolkas
parents: 52627
diff changeset
    33
    ArithM |
d6f2a7c196f7 added blast, force
smolkas
parents: 52627
diff changeset
    34
    BlastM
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    35
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    36
  (* legacy *)
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    37
  val By_Metis : facts -> byline
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
    38
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
    39
  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
    40
  val no_facts : facts
50268
5d6494332b0b added signature
smolkas
parents: 50267
diff changeset
    41
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    42
  val label_ord : label * label -> order
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    43
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    44
  val dummy_isar_step : isar_step
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    45
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51239
diff changeset
    46
  val string_of_label : label -> string
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    47
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
    48
  val fix_of_proof : isar_proof -> fix
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
    49
  val assms_of_proof : isar_proof -> assms
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
    50
  val steps_of_proof : isar_proof -> isar_step list
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
    51
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
    52
  val label_of_step : isar_step -> label option
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    53
  val subproofs_of_step : isar_step -> isar_proof 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
    54
  val byline_of_step : isar_step -> byline option
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    55
  val proof_method_of_step : isar_step -> proof_method 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
    56
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    57
  val fold_isar_step : (isar_step -> 'a -> 'a) -> isar_step -> 'a -> 'a
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    58
  val fold_isar_steps :
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    59
    (isar_step -> 'a -> 'a) -> isar_step list -> 'a -> 'a
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    60
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    61
  val map_isar_steps : (isar_step -> isar_step) -> isar_proof -> isar_proof
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    62
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    63
  val map_facts_of_byline : (facts -> facts) -> byline -> byline
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    64
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    65
  val add_proof_steps : isar_step list -> int -> int
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    66
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    67
  (** canonical proof labels: 1, 2, 3, ... in post traversal order **)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    68
  val canonical_label_ord : (label * label) -> order
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    69
  val relabel_proof_canonically : isar_proof -> isar_proof
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    70
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    71
  structure Canonical_Lbl_Tab : TABLE
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    72
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    73
end
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    74
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50269
diff changeset
    75
structure Sledgehammer_Proof : SLEDGEHAMMER_PROOF =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    76
struct
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    77
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    78
type label = string * int
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    79
type facts = label list * string list (* local & global facts *)
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    80
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
    81
datatype isar_qualifier = Show | Then
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    82
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
    83
datatype fix = Fix of (string * typ) list
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
    84
datatype assms = Assume of (label * term) list
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
    85
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    86
datatype isar_proof =
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
    87
  Proof of fix * assms * isar_step list
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
    88
and isar_step =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    89
  Let of term * term |
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    90
  (* for |fix|>0, this is an obtain step; step may contain subproofs *)
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
    91
  Prove of isar_qualifier list * fix * label * term * isar_proof list * byline
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    92
and byline =
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    93
  By of facts * proof_method
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    94
and proof_method =
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    95
  MetisM |
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    96
  SimpM |
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    97
  AutoM |
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
    98
  FastforceM |
52629
d6f2a7c196f7 added blast, force
smolkas
parents: 52627
diff changeset
    99
  ForceM |
d6f2a7c196f7 added blast, force
smolkas
parents: 52627
diff changeset
   100
  ArithM |
d6f2a7c196f7 added blast, force
smolkas
parents: 52627
diff changeset
   101
  BlastM
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   102
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   103
(* legacy *)
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   104
fun By_Metis facts = By (facts, MetisM)
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
   105
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
   106
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
   107
val no_facts = ([],[])
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   108
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   109
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
   110
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   111
val dummy_isar_step = Let (Term.dummy, Term.dummy)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   112
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51239
diff changeset
   113
fun string_of_label (s, num) = s ^ string_of_int num
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   114
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
   115
fun fix_of_proof (Proof (fix, _, _)) = fix
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
   116
fun assms_of_proof (Proof (_, assms, _)) = assms
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
   117
fun steps_of_proof (Proof (_, _, steps)) = steps
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
   118
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   119
fun label_of_step (Prove (_, _, l, _, _, _)) = SOME l
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
   120
  | label_of_step _ = NONE
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   121
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   122
fun subproofs_of_step (Prove (_, _, _, _, subproofs, _)) = SOME subproofs
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   123
  | subproofs_of_step _ = NONE
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   124
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   125
fun byline_of_step (Prove (_, _, _, _, _, byline)) = SOME byline
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
   126
  | byline_of_step _ = NONE
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
   127
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   128
fun proof_method_of_step (Prove (_, _, _, _, _, By(_, method))) = SOME method
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   129
  | proof_method_of_step _ = NONE
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   130
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   131
fun fold_isar_steps f = fold (fold_isar_step f)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   132
and fold_isar_step f step s =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   133
  fold (steps_of_proof #> fold_isar_steps f)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   134
       (these (subproofs_of_step step)) s
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   135
    |> f step
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   136
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   137
fun map_isar_steps f proof =
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   138
  let
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   139
    fun do_proof (Proof (fix, assms, steps)) =
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   140
      Proof (fix, assms, map do_step steps)
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   141
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   142
    and do_step (step as Let _) = f step
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   143
      | do_step (Prove (qs, xs, l, t, subproofs, by)) =
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   144
          let
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   145
            val subproofs = map do_proof subproofs
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   146
            val step = Prove (qs, xs, l, t, subproofs, by)
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   147
          in
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   148
            f step
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   149
          end
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   150
  in
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   151
    do_proof proof
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   152
  end
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   153
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   154
fun map_facts_of_byline f (By (facts, method)) = By (f facts, method)
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   155
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   156
val add_proof_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
   157
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   158
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   159
(** canonical proof labels: 1, 2, 3, ... in post traversal order **)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   160
52557
92ed2926596d made SML/NJ happy
smolkas
parents: 52556
diff changeset
   161
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
   162
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   163
structure Canonical_Lbl_Tab = Table(
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   164
  type key = label
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   165
  val ord = canonical_label_ord)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   166
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   167
fun relabel_proof_canonically proof =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   168
  let
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   169
    val lbl = pair ""
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   170
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   171
    fun next_label l (next, subst) =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   172
      (lbl next, (next + 1, (l, lbl next) :: subst))
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   173
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   174
    fun do_byline by (_, subst) =
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52590
diff changeset
   175
      by |> map_facts_of_byline (apfst (map (AList.lookup (op =) subst #> the)))
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   176
    handle Option.Option =>
52627
smolkas
parents: 52592
diff changeset
   177
      raise Fail "Sledgehammer_Proof: relabel_proof_canonically"
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   178
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   179
    fun do_assm (l, t) state =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   180
      let
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   181
        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
   182
      in ((l, t), state) end
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   183
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   184
    fun do_proof (Proof (fix, Assume assms, steps)) state =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   185
      let
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   186
        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
   187
        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
   188
      in
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   189
        (Proof (fix, Assume assms, steps), state)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   190
      end
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   191
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   192
    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
   193
      | do_step (Prove (qs, fix, l, t, subproofs, by)) state=
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   194
      let
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   195
        val by = do_byline by state
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   196
        val (subproofs, state) = fold_map do_proof subproofs state
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
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   199
        (Prove (qs, fix, l, t, subproofs, by), state)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   200
      end
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   201
  in
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   202
    fst (do_proof proof (0, []))
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   203
  end
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   204
50260
87ddf7eddfc9 simplified isar_qualifiers and qs merging
smolkas
parents: 50259
diff changeset
   205
end