src/Pure/search.ML
author wenzelm
Tue, 22 Jan 2019 19:36:17 +0100
changeset 69719 331ef175a112
parent 62919 9eb0359d6a77
child 81954 6f2bcdfa9a19
permissions -rw-r--r--
Backed out changeset 1bc422c08209 -- obsolete in AFP/5d11846ac6ab;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22360
diff changeset
     1
(*  Title:      Pure/search.ML
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
     2
    Author:     Lawrence C Paulson
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
     3
    Author:     Norbert Voelker, FernUniversitaet Hagen
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
     4
16179
fa7e70be26b0 header;
wenzelm
parents: 15574
diff changeset
     5
Search tacticals.
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
     6
*)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
     7
2672
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2143
diff changeset
     8
infix 1 THEN_MAYBE THEN_MAYBE';
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2143
diff changeset
     9
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    10
signature SEARCH =
32940
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    11
sig
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    12
  val DEPTH_FIRST: (thm -> bool) -> tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    13
  val has_fewer_prems: int -> thm -> bool
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    14
  val IF_UNSOLVED: tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    15
  val SOLVE: tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    16
  val THEN_MAYBE: tactic * tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    17
  val THEN_MAYBE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    18
  val DEPTH_SOLVE: tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    19
  val DEPTH_SOLVE_1: tactic -> tactic
38802
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 35408
diff changeset
    20
  val THEN_ITER_DEEPEN: int -> tactic -> (thm -> bool) -> (int -> tactic) -> tactic
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 35408
diff changeset
    21
  val ITER_DEEPEN: int -> (thm -> bool) -> (int -> tactic) -> tactic
32940
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    22
  val DEEPEN: int * int -> (int -> int -> tactic) -> int -> int -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    23
  val THEN_BEST_FIRST: tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    24
  val BEST_FIRST: (thm -> bool) * (thm -> int) -> tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    25
  val BREADTH_FIRST: (thm -> bool) -> tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    26
  val QUIET_BREADTH_FIRST: (thm -> bool) -> tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    27
  val THEN_ASTAR: tactic -> (thm -> bool) * (int -> thm -> int) -> tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    28
  val ASTAR: (thm -> bool) * (int -> thm -> int) -> tactic -> tactic
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    29
end;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    30
32940
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
    31
structure Search: SEARCH =
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    32
struct
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    33
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    34
(**** Depth-first search ****)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    35
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    36
(*Searches until "satp" reports proof tree as satisfied.
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    37
  Suppresses duplicate solutions to minimize search space.*)
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22360
diff changeset
    38
fun DEPTH_FIRST satp tac =
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    39
  let
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    40
    fun depth used [] = NONE
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    41
      | depth used (q :: qs) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    42
          (case Seq.pull q of
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    43
            NONE => depth used qs
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    44
          | SOME (st, stq) =>
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    45
              if satp st andalso not (member Thm.eq_thm used st) then
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    46
                SOME (st, Seq.make (fn() => depth (st :: used) (stq :: qs)))
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    47
              else depth used (tac st :: stq :: qs));
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
    48
  in fn st => Seq.make (fn () => depth [] [Seq.single st]) end;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    49
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    50
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    51
(*Predicate: Does the rule have fewer than n premises?*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
    52
fun has_fewer_prems n rule = Thm.nprems_of rule < n;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    53
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    54
(*Apply a tactic if subgoals remain, else do nothing.*)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    55
val IF_UNSOLVED = COND (has_fewer_prems 1) all_tac;
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    56
5754
98744e38ded1 added SOLVE tactical
oheimb
parents: 5693
diff changeset
    57
(*Force a tactic to solve its goal completely, otherwise fail *)
98744e38ded1 added SOLVE tactical
oheimb
parents: 5693
diff changeset
    58
fun SOLVE tac = tac THEN COND (has_fewer_prems 1) all_tac no_tac;
98744e38ded1 added SOLVE tactical
oheimb
parents: 5693
diff changeset
    59
2672
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2143
diff changeset
    60
(*Execute tac1, but only execute tac2 if there are at least as many subgoals
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2143
diff changeset
    61
  as before.  This ensures that tac2 is only applied to an outcome of tac1.*)
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22360
diff changeset
    62
fun (tac1 THEN_MAYBE tac2) st =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
    63
  (tac1 THEN COND (has_fewer_prems (Thm.nprems_of st)) all_tac tac2) st;
2672
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2143
diff changeset
    64
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2143
diff changeset
    65
fun (tac1 THEN_MAYBE' tac2) x = tac1 x THEN_MAYBE tac2 x;
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2143
diff changeset
    66
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    67
(*Tactical to reduce the number of premises by 1.
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    68
  If no subgoals then it must fail! *)
3538
ed9de44032e0 Removal of the tactical STATE
paulson
parents: 2869
diff changeset
    69
fun DEPTH_SOLVE_1 tac st = st |>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
    70
  (case Thm.nprems_of st of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
    71
    0 => no_tac
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
    72
  | n => DEPTH_FIRST (has_fewer_prems n) tac);
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    73
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    74
(*Uses depth-first search to solve ALL subgoals*)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    75
val DEPTH_SOLVE = DEPTH_FIRST (has_fewer_prems 1);
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    76
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    77
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    78
2869
acee08856cc9 DEEPEN now takes an upper bound for terminating searches
paulson
parents: 2672
diff changeset
    79
(**** Iterative deepening with pruning ****)
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    80
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    81
fun has_vars (Var _) = true
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    82
  | has_vars (Abs (_, _, t)) = has_vars t
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    83
  | has_vars (f $ t) = has_vars f orelse has_vars t
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    84
  | has_vars _ = false;
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    85
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    86
(*Counting of primitive inferences is APPROXIMATE, as the step tactic
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    87
  may perform >1 inference*)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    88
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
    89
(*Pruning of rigid ancestor to prevent backtracking*)
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22360
diff changeset
    90
fun prune (new as (k', np':int, rgd', stq), qs) =
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    91
  let
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    92
    fun prune_aux (qs, []) = new :: qs
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    93
      | prune_aux (qs, (k, np, rgd, q) :: rqs) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    94
          if np' + 1 = np andalso rgd then
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    95
             (*Use OLD k: zero-cost solution; see Stickel, p 365*)
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
    96
             (k, np', rgd', stq) :: qs
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    97
          else prune_aux ((k, np, rgd, q) :: qs, rqs)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    98
      fun take ([], rqs) = ([], rqs)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
    99
        | take (arg as ((k, np, rgd, stq) :: qs, rqs)) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   100
            if np' < np then take (qs, (k, np, rgd, stq) :: rqs) else arg;
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   101
  in prune_aux (take (qs, [])) end;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   102
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   103
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   104
(*Depth-first iterative deepening search for a state that satisfies satp
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   105
  tactic tac0 sets up the initial goal queue, while tac1 searches it.
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   106
  The solution sequence is redundant: the cutoff heuristic makes it impossible
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   107
  to suppress solutions arising from earlier searches, as the accumulated cost
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   108
  (k) can be wrong.*)
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   109
fun THEN_ITER_DEEPEN lim tac0 satp tac1 st =
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   110
  let
62919
wenzelm
parents: 62916
diff changeset
   111
    val counter = Unsynchronized.ref 0
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   112
    and tf = tac1 1
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   113
    and qs0 = tac0 st
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   114
     (*bnd = depth bound; inc = estimate of increment required next*)
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   115
    fun depth (bnd, inc) [] =
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   116
          if bnd > lim then NONE
22025
7c5896919eb8 Introduction of a bound on DEPTH_ITER, to prevent the meson method from running forever
paulson
parents: 20852
diff changeset
   117
          else
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   118
            (*larger increments make it run slower for the hard problems*)
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   119
            depth (bnd + inc, 10) [(0, 1, false, qs0)]
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   120
      | depth (bnd, inc) ((k, np, rgd, q) :: qs) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   121
          if k >= bnd then depth (bnd, inc) qs
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   122
          else
62919
wenzelm
parents: 62916
diff changeset
   123
           (case (Unsynchronized.inc counter; Seq.pull q) of
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   124
             NONE => depth (bnd, inc) qs
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   125
           | SOME (st, stq) =>
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   126
              if satp st then (*solution!*)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   127
                SOME(st, Seq.make (fn() => depth (bnd, inc) ((k, np, rgd, stq) :: qs)))
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   128
              else
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   129
                let
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   130
                  val np' = Thm.nprems_of st;
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   131
                  (*rgd' calculation assumes tactic operates on subgoal 1*)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   132
                  val rgd' = not (has_vars (hd (Thm.prems_of st)));
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   133
                  val k' = k + np' - np + 1;  (*difference in # of subgoals, +1*)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   134
                in
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   135
                  if k' + np' >= bnd then depth (bnd, Int.min (inc, k' + np' + 1 - bnd)) qs
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   136
                  else if np' < np (*solved a subgoal; prune rigid ancestors*)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   137
                  then depth (bnd, inc) (prune ((k', np', rgd', tf st), (k, np, rgd, stq) :: qs))
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   138
                  else depth (bnd, inc) ((k', np', rgd', tf st) :: (k, np, rgd, stq) :: qs)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   139
                end)
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   140
  in Seq.make (fn () => depth (0, 5) []) end;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   141
38802
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 35408
diff changeset
   142
fun ITER_DEEPEN lim = THEN_ITER_DEEPEN lim all_tac;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   143
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   144
2869
acee08856cc9 DEEPEN now takes an upper bound for terminating searches
paulson
parents: 2672
diff changeset
   145
(*Simple iterative deepening tactical.  It merely "deepens" any search tactic
acee08856cc9 DEEPEN now takes an upper bound for terminating searches
paulson
parents: 2672
diff changeset
   146
  using increment "inc" up to limit "lim". *)
38802
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 35408
diff changeset
   147
fun DEEPEN (inc, lim) tacf m i =
33380
cd6023a9a922 DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents: 33335
diff changeset
   148
  let
cd6023a9a922 DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents: 33335
diff changeset
   149
    fun dpn m st =
cd6023a9a922 DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents: 33335
diff changeset
   150
      st |>
cd6023a9a922 DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents: 33335
diff changeset
   151
       (if has_fewer_prems i st then no_tac
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   152
        else if m > lim then no_tac
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   153
        else tacf m i  ORELSE  dpn (m+inc))
2869
acee08856cc9 DEEPEN now takes an upper bound for terminating searches
paulson
parents: 2672
diff changeset
   154
  in  dpn m  end;
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22360
diff changeset
   155
32940
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   156
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   157
(*** Best-first search ***)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   158
32940
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   159
(*total ordering on theorems, allowing duplicates to be found*)
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   160
structure Thm_Heap = Heap
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   161
(
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   162
  type elem = int * thm;
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 55171
diff changeset
   163
  val ord = prod_ord int_ord (Term_Ord.term_ord o apply2 Thm.prop_of);
32940
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   164
);
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   165
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   166
(*For creating output sequence*)
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   167
fun some_of_list [] = NONE
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   168
  | some_of_list (x :: l) = SOME (x, Seq.make (fn () => some_of_list l));
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   169
9094
530e7a33b3dd now uses the heap data structure for BEST_FIRST
paulson
parents: 8149
diff changeset
   170
(*Check for and delete duplicate proof states*)
32940
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   171
fun delete_all_min prf heap =
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   172
  if Thm_Heap.is_empty heap then heap
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   173
  else if Thm.eq_thm (prf, #2 (Thm_Heap.min heap))
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   174
  then delete_all_min prf (Thm_Heap.delete_min heap)
51d450f278ce tuned signature;
wenzelm
parents: 32939
diff changeset
   175
  else heap;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   176
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   177
(*Best-first search for a state that satisfies satp (incl initial state)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   178
  Function sizef estimates size of problem remaining (smaller means better).
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   179
  tactic tac0 sets up the initial priority queue, while tac1 searches it. *)
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   180
fun THEN_BEST_FIRST tac0 (satp, sizef) tac =
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   181
  let
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   182
    fun pairsize th = (sizef th, th);
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   183
    fun bfs (news, nprf_heap) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   184
      (case List.partition satp news of
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   185
        ([], nonsats) => next (fold_rev Thm_Heap.insert (map pairsize nonsats) nprf_heap)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   186
      | (sats, _)  => some_of_list sats)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   187
    and next nprf_heap =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   188
      if Thm_Heap.is_empty nprf_heap then NONE
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   189
      else
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   190
        let
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   191
          val (n, prf) = Thm_Heap.min nprf_heap;
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   192
        in
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   193
          bfs (Seq.list_of (tac prf), delete_all_min prf (Thm_Heap.delete_min nprf_heap))
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   194
        end;
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   195
    fun btac st = bfs (Seq.list_of (tac0 st), Thm_Heap.empty)
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   196
  in fn st => Seq.make (fn () => btac st) end;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   197
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   198
(*Ordinary best-first search, with no initial tactic*)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   199
val BEST_FIRST = THEN_BEST_FIRST all_tac;
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   200
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22360
diff changeset
   201
(*Breadth-first search to satisfy satpred (including initial state)
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   202
  SLOW -- SHOULD NOT USE APPEND!*)
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   203
fun gen_BREADTH_FIRST message satpred (tac: tactic) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   204
  let
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   205
    val tacf = Seq.list_of o tac;
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   206
    fun bfs prfs =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   207
      (case List.partition satpred prfs of
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   208
        ([], []) => []
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   209
      | ([], nonsats) =>
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   210
          (message ("breadth=" ^ string_of_int (length nonsats));
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   211
           bfs (maps tacf nonsats))
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   212
      | (sats, _)  => sats);
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   213
  in fn st => Seq.of_list (bfs [st]) end;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   214
12262
11ff5f47df6e use tracing function for trace output;
wenzelm
parents: 9411
diff changeset
   215
val BREADTH_FIRST = gen_BREADTH_FIRST tracing;
5693
998af7667fa3 QUIET_BREADTH_FIRST;
wenzelm
parents: 4270
diff changeset
   216
val QUIET_BREADTH_FIRST = gen_BREADTH_FIRST (K ());
998af7667fa3 QUIET_BREADTH_FIRST;
wenzelm
parents: 4270
diff changeset
   217
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   218
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   219
(*
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   220
  Implementation of A*-like proof procedure by modification
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   221
  of the existing code for BEST_FIRST and best_tac so that the
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   222
  current level of search is taken into account.
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22360
diff changeset
   223
*)
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   224
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   225
(*Insertion into priority queue of states, marked with level*)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   226
fun insert_with_level (lnth: int * int * thm) [] = [lnth]
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   227
  | insert_with_level (l, m, th) ((l', n, th') :: nths) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   228
      if  n < m then (l', n, th') :: insert_with_level (l, m, th) nths
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   229
      else if n = m andalso Thm.eq_thm (th, th') then (l', n, th') :: nths
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   230
      else (l, m, th) :: (l', n, th') :: nths;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   231
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   232
(*For creating output sequence*)
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   233
fun some_of_list [] = NONE
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   234
  | some_of_list (x :: xs) = SOME (x, Seq.make (fn () => some_of_list xs));
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   235
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   236
fun THEN_ASTAR tac0 (satp, costf) tac =
60940
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   237
  let
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   238
    fun bfs (news, nprfs, level) =
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   239
      let fun cost thm = (level, costf level thm, thm) in
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   240
        (case List.partition satp news of
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   241
          ([], nonsats) => next (fold_rev (insert_with_level o cost) nonsats nprfs)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   242
        | (sats, _) => some_of_list sats)
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   243
      end
4c108cce6b35 tuned whitespace;
wenzelm
parents: 59582
diff changeset
   244
    and next [] = NONE
62916
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   245
      | next ((level, n, prf) :: nprfs) = bfs (Seq.list_of (tac prf), nprfs, level + 1)
621afc4607ec eliminated ancient TTY-based Tactical.tracify and related global references;
wenzelm
parents: 60940
diff changeset
   246
  in fn st => Seq.make (fn () => bfs (Seq.list_of (tac0 st), [], 0)) end;
1588
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   247
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   248
(*Ordinary ASTAR, with no initial tactic*)
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   249
val ASTAR = THEN_ASTAR all_tac;
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   250
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   251
end;
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   252
fff3738830f5 New file containing search tacticals
paulson
parents:
diff changeset
   253
open Search;