| 18120 |      1 | (*  Title:      Pure/old_goals.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1993  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Old-style goal stack package.  The goal stack initially holds a dummy
 | 
|  |      7 | proof, and can never become empty.  Each goal stack consists of a list
 | 
|  |      8 | of levels.  The undo list is a list of goal stacks.  Finally, there
 | 
|  |      9 | may be a stack of pending proofs.
 | 
|  |     10 | *)
 | 
|  |     11 | 
 | 
|  |     12 | signature GOALS =
 | 
|  |     13 | sig
 | 
|  |     14 |   val premises: unit -> thm list
 | 
|  |     15 |   val prove_goal: theory -> string -> (thm list -> tactic list) -> thm
 | 
|  |     16 |   val prove_goalw: theory -> thm list -> string -> (thm list -> tactic list) -> thm
 | 
|  |     17 |   val disable_pr: unit -> unit
 | 
|  |     18 |   val enable_pr: unit -> unit
 | 
|  |     19 |   val topthm: unit -> thm
 | 
|  |     20 |   val result: unit -> thm
 | 
|  |     21 |   val uresult: unit -> thm
 | 
|  |     22 |   val getgoal: int -> term
 | 
|  |     23 |   val gethyps: int -> thm list
 | 
|  |     24 |   val prlev: int -> unit
 | 
|  |     25 |   val pr: unit -> unit
 | 
|  |     26 |   val prlim: int -> unit
 | 
|  |     27 |   val goal: theory -> string -> thm list
 | 
|  |     28 |   val goalw: theory -> thm list -> string -> thm list
 | 
|  |     29 |   val Goal: string -> thm list
 | 
|  |     30 |   val Goalw: thm list -> string -> thm list
 | 
|  |     31 |   val by: tactic -> unit
 | 
|  |     32 |   val back: unit -> unit
 | 
|  |     33 |   val choplev: int -> unit
 | 
|  |     34 |   val undo: unit -> unit
 | 
|  |     35 |   val bind_thm: string * thm -> unit
 | 
|  |     36 |   val bind_thms: string * thm list -> unit
 | 
|  |     37 |   val qed: string -> unit
 | 
|  |     38 |   val qed_goal: string -> theory -> string -> (thm list -> tactic list) -> unit
 | 
|  |     39 |   val qed_goalw: string -> theory -> thm list -> string
 | 
|  |     40 |     -> (thm list -> tactic list) -> unit
 | 
|  |     41 |   val qed_spec_mp: string -> unit
 | 
|  |     42 |   val qed_goal_spec_mp: string -> theory -> string -> (thm list -> tactic list) -> unit
 | 
|  |     43 |   val qed_goalw_spec_mp: string -> theory -> thm list -> string
 | 
|  |     44 |     -> (thm list -> tactic list) -> unit
 | 
|  |     45 |   val no_qed: unit -> unit
 | 
|  |     46 | end;
 | 
|  |     47 | 
 | 
|  |     48 | signature OLD_GOALS =
 | 
|  |     49 | sig
 | 
|  |     50 |   include GOALS
 | 
|  |     51 |   val legacy: bool ref
 | 
|  |     52 |   type proof
 | 
| 19053 |     53 |   val chop: unit -> unit
 | 
| 18120 |     54 |   val reset_goals: unit -> unit
 | 
|  |     55 |   val result_error_fn: (thm -> string -> thm) ref
 | 
|  |     56 |   val print_sign_exn: theory -> exn -> 'a
 | 
|  |     57 |   val prove_goalw_cterm: thm list->cterm->(thm list->tactic list)->thm
 | 
|  |     58 |   val prove_goalw_cterm_nocheck: thm list->cterm->(thm list->tactic list)->thm
 | 
|  |     59 |   val quick_and_dirty_prove_goalw_cterm: theory -> thm list -> cterm
 | 
|  |     60 |     -> (thm list -> tactic list) -> thm
 | 
|  |     61 |   val print_exn: exn -> 'a
 | 
|  |     62 |   val filter_goal: (term*term->bool) -> thm list -> int -> thm list
 | 
|  |     63 |   val goalw_cterm: thm list -> cterm -> thm list
 | 
|  |     64 |   val simple_prove_goal_cterm: cterm->(thm list->tactic list)->thm
 | 
|  |     65 |   val byev: tactic list -> unit
 | 
|  |     66 |   val save_proof: unit -> proof
 | 
|  |     67 |   val restore_proof: proof -> thm list
 | 
|  |     68 |   val push_proof: unit -> unit
 | 
|  |     69 |   val pop_proof: unit -> thm list
 | 
|  |     70 |   val rotate_proof: unit -> thm list
 | 
|  |     71 |   val bws: thm list -> unit
 | 
|  |     72 |   val bw: thm -> unit
 | 
|  |     73 |   val brs: thm list -> int -> unit
 | 
|  |     74 |   val br: thm -> int -> unit
 | 
|  |     75 |   val bes: thm list -> int -> unit
 | 
|  |     76 |   val be: thm -> int -> unit
 | 
|  |     77 |   val bds: thm list -> int -> unit
 | 
|  |     78 |   val bd: thm -> int -> unit
 | 
|  |     79 |   val ba: int -> unit
 | 
|  |     80 |   val ren: string -> int -> unit
 | 
|  |     81 |   val frs: thm list -> unit
 | 
|  |     82 |   val fr: thm -> unit
 | 
|  |     83 |   val fes: thm list -> unit
 | 
|  |     84 |   val fe: thm -> unit
 | 
|  |     85 |   val fds: thm list -> unit
 | 
|  |     86 |   val fd: thm -> unit
 | 
|  |     87 |   val fa: unit -> unit
 | 
|  |     88 | end;
 | 
|  |     89 | 
 | 
|  |     90 | structure OldGoals: OLD_GOALS =
 | 
|  |     91 | struct
 | 
|  |     92 | 
 | 
|  |     93 | val legacy = ref false;
 | 
|  |     94 | fun warn_obsolete () = if ! legacy then () else warning "Obsolete goal command encountered";
 | 
|  |     95 | 
 | 
|  |     96 | 
 | 
|  |     97 | (*** Goal package ***)
 | 
|  |     98 | 
 | 
|  |     99 | (*Each level of goal stack includes a proof state and alternative states,
 | 
|  |    100 |   the output of the tactic applied to the preceeding level.  *)
 | 
|  |    101 | type gstack = (thm * thm Seq.seq) list;
 | 
|  |    102 | 
 | 
|  |    103 | datatype proof = Proof of gstack list * thm list * (bool*thm->thm);
 | 
|  |    104 | 
 | 
|  |    105 | 
 | 
|  |    106 | (*** References ***)
 | 
|  |    107 | 
 | 
|  |    108 | (*Current assumption list -- set by "goal".*)
 | 
|  |    109 | val curr_prems = ref([] : thm list);
 | 
|  |    110 | 
 | 
|  |    111 | (*Return assumption list -- useful if you didn't save "goal"'s result. *)
 | 
|  |    112 | fun premises() = !curr_prems;
 | 
|  |    113 | 
 | 
|  |    114 | (*Current result maker -- set by "goal", used by "result".  *)
 | 
|  |    115 | fun init_mkresult _ = error "No goal has been supplied in subgoal module";
 | 
|  |    116 | val curr_mkresult = ref (init_mkresult: bool*thm->thm);
 | 
|  |    117 | 
 | 
|  |    118 | val dummy = Thm.trivial (read_cterm ProtoPure.thy ("PROP No_goal_has_been_supplied", propT));
 | 
|  |    119 | 
 | 
|  |    120 | (*List of previous goal stacks, for the undo operation.  Set by setstate.
 | 
|  |    121 |   A list of lists!*)
 | 
|  |    122 | val undo_list = ref([[(dummy, Seq.empty)]] : gstack list);
 | 
|  |    123 | 
 | 
|  |    124 | (* Stack of proof attempts *)
 | 
|  |    125 | val proofstack = ref([]: proof list);
 | 
|  |    126 | 
 | 
|  |    127 | (*reset all refs*)
 | 
|  |    128 | fun reset_goals () =
 | 
|  |    129 |   (curr_prems := []; curr_mkresult := init_mkresult;
 | 
|  |    130 |     undo_list := [[(dummy, Seq.empty)]]);
 | 
|  |    131 | 
 | 
|  |    132 | 
 | 
|  |    133 | (*** Setting up goal-directed proof ***)
 | 
|  |    134 | 
 | 
|  |    135 | (*Generates the list of new theories when the proof state's theory changes*)
 | 
|  |    136 | fun thy_error (thy,thy') =
 | 
|  |    137 |   let val names = Context.names_of thy' \\ Context.names_of thy
 | 
|  |    138 |   in  case names of
 | 
|  |    139 |         [name] => "\nNew theory: " ^ name
 | 
|  |    140 |       | _       => "\nNew theories: " ^ space_implode ", " names
 | 
|  |    141 |   end;
 | 
|  |    142 | 
 | 
|  |    143 | (*Default action is to print an error message; could be suppressed for
 | 
|  |    144 |   special applications.*)
 | 
|  |    145 | fun result_error_default state msg : thm =
 | 
|  |    146 |   Pretty.str "Bad final proof state:" :: Display.pretty_goals (!goals_limit) state @
 | 
|  |    147 |     [Pretty.str msg, Pretty.str "Proof failed!"] |> Pretty.chunks |> Pretty.string_of |> error;
 | 
|  |    148 | 
 | 
|  |    149 | val result_error_fn = ref result_error_default;
 | 
|  |    150 | 
 | 
|  |    151 | 
 | 
|  |    152 | (*Common treatment of "goal" and "prove_goal":
 | 
|  |    153 |   Return assumptions, initial proof state, and function to make result.
 | 
|  |    154 |   "atomic" indicates if the goal should be wrapped up in the function
 | 
|  |    155 |   "Goal::prop=>prop" to avoid assumptions being returned separately.
 | 
|  |    156 | *)
 | 
|  |    157 | fun prepare_proof atomic rths chorn =
 | 
|  |    158 |   let
 | 
|  |    159 |       val _ = warn_obsolete ();
 | 
|  |    160 |       val {thy, t=horn,...} = rep_cterm chorn;
 | 
|  |    161 |       val _ = Term.no_dummy_patterns horn handle TERM (msg, _) => error msg;
 | 
|  |    162 |       val (As, B) = Logic.strip_horn horn;
 | 
|  |    163 |       val atoms = atomic andalso
 | 
|  |    164 |             forall (fn t => not(Logic.is_implies t orelse Logic.is_all t)) As;
 | 
|  |    165 |       val (As,B) = if atoms then ([],horn) else (As,B);
 | 
|  |    166 |       val cAs = map (cterm_of thy) As;
 | 
|  |    167 |       val prems = map (rewrite_rule rths o forall_elim_vars 0 o Thm.assume) cAs;
 | 
|  |    168 |       val cB = cterm_of thy B;
 | 
|  |    169 |       val st0 = let val st = Goal.init cB |> fold Thm.weaken cAs
 | 
|  |    170 |                 in  rewrite_goals_rule rths st end
 | 
|  |    171 |       (*discharges assumptions from state in the order they appear in goal;
 | 
|  |    172 |         checks (if requested) that resulting theorem is equivalent to goal. *)
 | 
|  |    173 |       fun mkresult (check,state) =
 | 
|  |    174 |         let val state = Seq.hd (flexflex_rule state)
 | 
|  |    175 |                         handle THM _ => state   (*smash flexflex pairs*)
 | 
|  |    176 |             val ngoals = nprems_of state
 | 
|  |    177 |             val ath = implies_intr_list cAs state
 | 
|  |    178 |             val th = Goal.conclude ath
 | 
|  |    179 |             val {hyps,prop,thy=thy',...} = rep_thm th
 | 
|  |    180 |             val final_th = standard th
 | 
|  |    181 |         in  if not check then final_th
 | 
|  |    182 |             else if not (eq_thy(thy,thy')) then !result_error_fn state
 | 
|  |    183 |                 ("Theory of proof state has changed!" ^
 | 
|  |    184 |                  thy_error (thy,thy'))
 | 
|  |    185 |             else if ngoals>0 then !result_error_fn state
 | 
|  |    186 |                 (string_of_int ngoals ^ " unsolved goals!")
 | 
|  |    187 |             else if not (null hyps) then !result_error_fn state
 | 
|  |    188 |                 ("Additional hypotheses:\n" ^
 | 
|  |    189 |                  cat_lines (map (Sign.string_of_term thy) hyps))
 | 
|  |    190 |             else if Pattern.matches thy
 | 
|  |    191 |                                     (Envir.beta_norm (term_of chorn), Envir.beta_norm prop)
 | 
|  |    192 |                  then final_th
 | 
|  |    193 |             else  !result_error_fn state "proved a different theorem"
 | 
|  |    194 |         end
 | 
|  |    195 |   in
 | 
|  |    196 |      if eq_thy(thy, Thm.theory_of_thm st0)
 | 
|  |    197 |      then (prems, st0, mkresult)
 | 
|  |    198 |      else error ("Definitions would change the proof state's theory" ^
 | 
|  |    199 |                  thy_error (thy, Thm.theory_of_thm st0))
 | 
|  |    200 |   end
 | 
|  |    201 |   handle THM(s,_,_) => error("prepare_proof: exception THM was raised!\n" ^ s);
 | 
|  |    202 | 
 | 
|  |    203 | (*Prints exceptions readably to users*)
 | 
|  |    204 | fun print_sign_exn_unit thy e =
 | 
|  |    205 |   case e of
 | 
|  |    206 |      THM (msg,i,thms) =>
 | 
|  |    207 |          (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg);
 | 
|  |    208 |           List.app print_thm thms)
 | 
|  |    209 |    | THEORY (msg,thys) =>
 | 
|  |    210 |          (writeln ("Exception THEORY raised:\n" ^ msg);
 | 
|  |    211 |           List.app (writeln o Context.str_of_thy) thys)
 | 
|  |    212 |    | TERM (msg,ts) =>
 | 
|  |    213 |          (writeln ("Exception TERM raised:\n" ^ msg);
 | 
|  |    214 |           List.app (writeln o Sign.string_of_term thy) ts)
 | 
|  |    215 |    | TYPE (msg,Ts,ts) =>
 | 
|  |    216 |          (writeln ("Exception TYPE raised:\n" ^ msg);
 | 
|  |    217 |           List.app (writeln o Sign.string_of_typ thy) Ts;
 | 
|  |    218 |           List.app (writeln o Sign.string_of_term thy) ts)
 | 
|  |    219 |    | e => raise e;
 | 
|  |    220 | 
 | 
|  |    221 | (*Prints an exception, then fails*)
 | 
| 18678 |    222 | fun print_sign_exn thy e = (print_sign_exn_unit thy e; raise ERROR "");
 | 
| 18120 |    223 | 
 | 
|  |    224 | (** the prove_goal.... commands
 | 
|  |    225 |     Prove theorem using the listed tactics; check it has the specified form.
 | 
|  |    226 |     Augment theory with all type assignments of goal.
 | 
|  |    227 |     Syntax is similar to "goal" command for easy keyboard use. **)
 | 
|  |    228 | 
 | 
|  |    229 | (*Version taking the goal as a cterm*)
 | 
|  |    230 | fun prove_goalw_cterm_general check rths chorn tacsf =
 | 
|  |    231 |   let val (prems, st0, mkresult) = prepare_proof false rths chorn
 | 
|  |    232 |       val tac = EVERY (tacsf prems)
 | 
|  |    233 |       fun statef() =
 | 
|  |    234 |           (case Seq.pull (tac st0) of
 | 
|  |    235 |                SOME(st,_) => st
 | 
|  |    236 |              | _ => error ("prove_goal: tactic failed"))
 | 
|  |    237 |   in  mkresult (check, cond_timeit (!Output.timing) statef)  end
 | 
|  |    238 |   handle e => (print_sign_exn_unit (#thy (rep_cterm chorn)) e;
 | 
|  |    239 |                writeln ("The exception above was raised for\n" ^
 | 
|  |    240 |                       Display.string_of_cterm chorn); raise e);
 | 
|  |    241 | 
 | 
|  |    242 | (*Two variants: one checking the result, one not.
 | 
|  |    243 |   Neither prints runtime messages: they are for internal packages.*)
 | 
|  |    244 | fun prove_goalw_cterm rths chorn =
 | 
|  |    245 |         setmp Output.timing false (prove_goalw_cterm_general true rths chorn)
 | 
|  |    246 | and prove_goalw_cterm_nocheck rths chorn =
 | 
|  |    247 |         setmp Output.timing false (prove_goalw_cterm_general false rths chorn);
 | 
|  |    248 | 
 | 
|  |    249 | 
 | 
|  |    250 | (*Version taking the goal as a string*)
 | 
|  |    251 | fun prove_goalw thy rths agoal tacsf =
 | 
|  |    252 |   let val chorn = read_cterm thy (agoal, propT)
 | 
|  |    253 |   in prove_goalw_cterm_general true rths chorn tacsf end
 | 
| 18678 |    254 |   handle ERROR msg => cat_error msg (*from read_cterm?*)
 | 
| 18120 |    255 |                 ("The error(s) above occurred for " ^ quote agoal);
 | 
|  |    256 | 
 | 
|  |    257 | (*String version with no meta-rewrite-rules*)
 | 
|  |    258 | fun prove_goal thy = prove_goalw thy [];
 | 
|  |    259 | 
 | 
|  |    260 | (*quick and dirty version (conditional)*)
 | 
|  |    261 | fun quick_and_dirty_prove_goalw_cterm thy rews ct tacs =
 | 
|  |    262 |   prove_goalw_cterm rews ct
 | 
|  |    263 |     (if ! quick_and_dirty then (K [SkipProof.cheat_tac thy]) else tacs);
 | 
|  |    264 | 
 | 
|  |    265 | 
 | 
|  |    266 | (*** Commands etc ***)
 | 
|  |    267 | 
 | 
|  |    268 | (*Return the current goal stack, if any, from undo_list*)
 | 
|  |    269 | fun getstate() : gstack = case !undo_list of
 | 
|  |    270 |       []   => error"No current state in subgoal module"
 | 
|  |    271 |     | x::_ => x;
 | 
|  |    272 | 
 | 
|  |    273 | (*Pops the given goal stack*)
 | 
|  |    274 | fun pop [] = error"Cannot go back past the beginning of the proof!"
 | 
|  |    275 |   | pop (pair::pairs) = (pair,pairs);
 | 
|  |    276 | 
 | 
|  |    277 | 
 | 
|  |    278 | (* Print a level of the goal stack -- subject to quiet mode *)
 | 
|  |    279 | 
 | 
|  |    280 | val quiet = ref false;
 | 
|  |    281 | fun disable_pr () = quiet := true;
 | 
|  |    282 | fun enable_pr () = quiet := false;
 | 
|  |    283 | 
 | 
|  |    284 | fun print_top ((th, _), pairs) =
 | 
|  |    285 |   if ! quiet then ()
 | 
|  |    286 |   else ! Display.print_current_goals_fn (length pairs) (! goals_limit) th;
 | 
|  |    287 | 
 | 
|  |    288 | (*Printing can raise exceptions, so the assignment occurs last.
 | 
|  |    289 |   Can do   setstate[(st,Seq.empty)]  to set st as the state.  *)
 | 
|  |    290 | fun setstate newgoals =
 | 
|  |    291 |   (print_top (pop newgoals);  undo_list := newgoals :: !undo_list);
 | 
|  |    292 | 
 | 
|  |    293 | (*Given a proof state transformation, return a command that updates
 | 
|  |    294 |     the goal stack*)
 | 
|  |    295 | fun make_command com = setstate (com (pop (getstate())));
 | 
|  |    296 | 
 | 
|  |    297 | (*Apply a function on proof states to the current goal stack*)
 | 
|  |    298 | fun apply_fun f = f (pop(getstate()));
 | 
|  |    299 | 
 | 
|  |    300 | (*Return the top theorem, representing the proof state*)
 | 
|  |    301 | fun topthm () = apply_fun  (fn ((th,_), _) => th);
 | 
|  |    302 | 
 | 
|  |    303 | (*Return the final result.  *)
 | 
|  |    304 | fun result () = !curr_mkresult (true, topthm());
 | 
|  |    305 | 
 | 
|  |    306 | (*Return the result UNCHECKED that it equals the goal -- for synthesis,
 | 
|  |    307 |   answer extraction, or other instantiation of Vars *)
 | 
|  |    308 | fun uresult () = !curr_mkresult (false, topthm());
 | 
|  |    309 | 
 | 
|  |    310 | (*Get subgoal i from goal stack*)
 | 
|  |    311 | fun getgoal i = Logic.get_goal (prop_of (topthm())) i;
 | 
|  |    312 | 
 | 
|  |    313 | (*Return subgoal i's hypotheses as meta-level assumptions.
 | 
|  |    314 |   For debugging uses of METAHYPS*)
 | 
|  |    315 | local exception GETHYPS of thm list
 | 
|  |    316 | in
 | 
|  |    317 | fun gethyps i =
 | 
|  |    318 |     (METAHYPS (fn hyps => raise (GETHYPS hyps)) i (topthm());  [])
 | 
|  |    319 |     handle GETHYPS hyps => hyps
 | 
|  |    320 | end;
 | 
|  |    321 | 
 | 
|  |    322 | (*Prints exceptions nicely at top level;
 | 
|  |    323 |   raises the exception in order to have a polymorphic type!*)
 | 
|  |    324 | fun print_exn e = (print_sign_exn_unit (Thm.theory_of_thm (topthm())) e;  raise e);
 | 
|  |    325 | 
 | 
|  |    326 | (*Which thms could apply to goal i? (debugs tactics involving filter_thms) *)
 | 
|  |    327 | fun filter_goal could ths i = filter_thms could (999, getgoal i, ths);
 | 
|  |    328 | 
 | 
|  |    329 | (*For inspecting earlier levels of the backward proof*)
 | 
|  |    330 | fun chop_level n (pair,pairs) =
 | 
|  |    331 |   let val level = length pairs
 | 
|  |    332 |   in  if n<0 andalso ~n <= level
 | 
|  |    333 |       then  List.drop (pair::pairs, ~n)
 | 
|  |    334 |       else if 0<=n andalso n<= level
 | 
|  |    335 |       then  List.drop (pair::pairs, level - n)
 | 
|  |    336 |       else  error ("Level number must lie between 0 and " ^
 | 
|  |    337 |                    string_of_int level)
 | 
|  |    338 |   end;
 | 
|  |    339 | 
 | 
|  |    340 | (*Print the given level of the proof; prlev ~1 prints previous level*)
 | 
|  |    341 | fun prlev n = (enable_pr (); apply_fun (print_top o pop o (chop_level n)));
 | 
|  |    342 | fun pr () = (enable_pr (); apply_fun print_top);
 | 
|  |    343 | 
 | 
|  |    344 | (*Set goals_limit and print again*)
 | 
|  |    345 | fun prlim n = (goals_limit:=n; pr());
 | 
|  |    346 | 
 | 
|  |    347 | (** the goal.... commands
 | 
|  |    348 |     Read main goal.  Set global variables curr_prems, curr_mkresult.
 | 
|  |    349 |     Initial subgoal and premises are rewritten using rths. **)
 | 
|  |    350 | 
 | 
|  |    351 | (*Version taking the goal as a cterm; if you have a term t and theory thy, use
 | 
|  |    352 |     goalw_cterm rths (cterm_of thy t);      *)
 | 
|  |    353 | fun agoalw_cterm atomic rths chorn =
 | 
|  |    354 |   let val (prems, st0, mkresult) = prepare_proof atomic rths chorn
 | 
|  |    355 |   in  undo_list := [];
 | 
|  |    356 |       setstate [ (st0, Seq.empty) ];
 | 
|  |    357 |       curr_prems := prems;
 | 
|  |    358 |       curr_mkresult := mkresult;
 | 
|  |    359 |       prems
 | 
|  |    360 |   end;
 | 
|  |    361 | 
 | 
|  |    362 | val goalw_cterm = agoalw_cterm false;
 | 
|  |    363 | 
 | 
|  |    364 | (*Version taking the goal as a string*)
 | 
|  |    365 | fun agoalw atomic thy rths agoal =
 | 
|  |    366 |     agoalw_cterm atomic rths (read_cterm thy (agoal, propT))
 | 
| 18678 |    367 |     handle ERROR msg => cat_error msg (*from type_assign, etc via prepare_proof*)
 | 
| 18120 |    368 |         ("The error(s) above occurred for " ^ quote agoal);
 | 
|  |    369 | 
 | 
|  |    370 | val goalw = agoalw false;
 | 
|  |    371 | fun goal thy = goalw thy [];
 | 
|  |    372 | 
 | 
|  |    373 | (*now the versions that wrap the goal up in `Goal' to make it atomic*)
 | 
|  |    374 | fun Goalw thms s = agoalw true (Context.the_context ()) thms s;
 | 
|  |    375 | val Goal = Goalw [];
 | 
|  |    376 | 
 | 
|  |    377 | (*simple version with minimal amount of checking and postprocessing*)
 | 
|  |    378 | fun simple_prove_goal_cterm G f =
 | 
|  |    379 |   let
 | 
|  |    380 |     val _ = warn_obsolete ();
 | 
|  |    381 |     val As = Drule.strip_imp_prems G;
 | 
|  |    382 |     val B = Drule.strip_imp_concl G;
 | 
|  |    383 |     val asms = map (Goal.norm_hhf o Thm.assume) As;
 | 
|  |    384 |     fun check NONE = error "prove_goal: tactic failed"
 | 
|  |    385 |       | check (SOME (thm, _)) = (case nprems_of thm of
 | 
|  |    386 |             0 => thm
 | 
|  |    387 |           | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
 | 
|  |    388 |   in
 | 
|  |    389 |     standard (implies_intr_list As
 | 
|  |    390 |       (check (Seq.pull (EVERY (f asms) (trivial B)))))
 | 
|  |    391 |   end;
 | 
|  |    392 | 
 | 
|  |    393 | 
 | 
|  |    394 | (*Proof step "by" the given tactic -- apply tactic to the proof state*)
 | 
|  |    395 | fun by_com tac ((th,ths), pairs) : gstack =
 | 
|  |    396 |   (case  Seq.pull(tac th)  of
 | 
|  |    397 |      NONE      => error"by: tactic failed"
 | 
|  |    398 |    | SOME(th2,ths2) =>
 | 
|  |    399 |        (if eq_thm(th,th2)
 | 
|  |    400 |           then warning "Warning: same as previous level"
 | 
|  |    401 |           else if eq_thm_thy(th,th2) then ()
 | 
|  |    402 |           else warning ("Warning: theory of proof state has changed" ^
 | 
|  |    403 |                        thy_error (Thm.theory_of_thm th, Thm.theory_of_thm th2));
 | 
|  |    404 |        ((th2,ths2)::(th,ths)::pairs)));
 | 
|  |    405 | 
 | 
|  |    406 | fun by tac = cond_timeit (!Output.timing)
 | 
|  |    407 |     (fn() => make_command (by_com tac));
 | 
|  |    408 | 
 | 
|  |    409 | (* byev[tac1,...,tacn] applies tac1 THEN ... THEN tacn.
 | 
|  |    410 |    Good for debugging proofs involving prove_goal.*)
 | 
|  |    411 | val byev = by o EVERY;
 | 
|  |    412 | 
 | 
|  |    413 | 
 | 
|  |    414 | (*Backtracking means find an alternative result from a tactic.
 | 
|  |    415 |   If none at this level, try earlier levels*)
 | 
|  |    416 | fun backtrack [] = error"back: no alternatives"
 | 
|  |    417 |   | backtrack ((th,thstr) :: pairs) =
 | 
|  |    418 |      (case Seq.pull thstr of
 | 
|  |    419 |           NONE      => (writeln"Going back a level..."; backtrack pairs)
 | 
|  |    420 |         | SOME(th2,thstr2) =>
 | 
|  |    421 |            (if eq_thm(th,th2)
 | 
|  |    422 |               then warning "Warning: same as previous choice at this level"
 | 
|  |    423 |               else if eq_thm_thy(th,th2) then ()
 | 
|  |    424 |               else warning "Warning: theory of proof state has changed";
 | 
|  |    425 |             (th2,thstr2)::pairs));
 | 
|  |    426 | 
 | 
|  |    427 | fun back() = setstate (backtrack (getstate()));
 | 
|  |    428 | 
 | 
|  |    429 | (*Chop back to previous level of the proof*)
 | 
|  |    430 | fun choplev n = make_command (chop_level n);
 | 
|  |    431 | 
 | 
|  |    432 | (*Chopping back the goal stack*)
 | 
|  |    433 | fun chop () = make_command (fn (_,pairs) => pairs);
 | 
|  |    434 | 
 | 
|  |    435 | (*Restore the previous proof state;  discard current state. *)
 | 
|  |    436 | fun undo() = case !undo_list of
 | 
|  |    437 |       [] => error"No proof state"
 | 
|  |    438 |     | [_] => error"Already at initial state"
 | 
|  |    439 |     | _::newundo =>  (undo_list := newundo;  pr()) ;
 | 
|  |    440 | 
 | 
|  |    441 | 
 | 
|  |    442 | (*** Managing the proof stack ***)
 | 
|  |    443 | 
 | 
|  |    444 | fun save_proof() = Proof(!undo_list, !curr_prems, !curr_mkresult);
 | 
|  |    445 | 
 | 
|  |    446 | fun restore_proof(Proof(ul,prems,mk)) =
 | 
|  |    447 |  (undo_list:= ul;  curr_prems:= prems;  curr_mkresult := mk;  prems);
 | 
|  |    448 | 
 | 
|  |    449 | 
 | 
|  |    450 | fun top_proof() = case !proofstack of
 | 
|  |    451 |         [] => error("Stack of proof attempts is empty!")
 | 
|  |    452 |     | p::ps => (p,ps);
 | 
|  |    453 | 
 | 
|  |    454 | (*  push a copy of the current proof state on to the stack *)
 | 
|  |    455 | fun push_proof() = (proofstack := (save_proof() :: !proofstack));
 | 
|  |    456 | 
 | 
|  |    457 | (* discard the top proof state of the stack *)
 | 
|  |    458 | fun pop_proof() =
 | 
|  |    459 |   let val (p,ps) = top_proof()
 | 
|  |    460 |       val prems = restore_proof p
 | 
|  |    461 |   in proofstack := ps;  pr();  prems end;
 | 
|  |    462 | 
 | 
|  |    463 | (* rotate the stack so that the top element goes to the bottom *)
 | 
|  |    464 | fun rotate_proof() = let val (p,ps) = top_proof()
 | 
|  |    465 |                     in proofstack := ps@[save_proof()];
 | 
|  |    466 |                        restore_proof p;
 | 
|  |    467 |                        pr();
 | 
|  |    468 |                        !curr_prems
 | 
|  |    469 |                     end;
 | 
|  |    470 | 
 | 
|  |    471 | 
 | 
|  |    472 | (** Shortcuts for commonly-used tactics **)
 | 
|  |    473 | 
 | 
|  |    474 | fun bws rls = by (rewrite_goals_tac rls);
 | 
|  |    475 | fun bw rl = bws [rl];
 | 
|  |    476 | 
 | 
|  |    477 | fun brs rls i = by (resolve_tac rls i);
 | 
|  |    478 | fun br rl = brs [rl];
 | 
|  |    479 | 
 | 
|  |    480 | fun bes rls i = by (eresolve_tac rls i);
 | 
|  |    481 | fun be rl = bes [rl];
 | 
|  |    482 | 
 | 
|  |    483 | fun bds rls i = by (dresolve_tac rls i);
 | 
|  |    484 | fun bd rl = bds [rl];
 | 
|  |    485 | 
 | 
|  |    486 | fun ba i = by (assume_tac i);
 | 
|  |    487 | 
 | 
|  |    488 | fun ren str i = by (rename_tac str i);
 | 
|  |    489 | 
 | 
|  |    490 | (** Shortcuts to work on the first applicable subgoal **)
 | 
|  |    491 | 
 | 
|  |    492 | fun frs rls = by (FIRSTGOAL (trace_goalno_tac (resolve_tac rls)));
 | 
|  |    493 | fun fr rl = frs [rl];
 | 
|  |    494 | 
 | 
|  |    495 | fun fes rls = by (FIRSTGOAL (trace_goalno_tac (eresolve_tac rls)));
 | 
|  |    496 | fun fe rl = fes [rl];
 | 
|  |    497 | 
 | 
|  |    498 | fun fds rls = by (FIRSTGOAL (trace_goalno_tac (dresolve_tac rls)));
 | 
|  |    499 | fun fd rl = fds [rl];
 | 
|  |    500 | 
 | 
|  |    501 | fun fa() = by (FIRSTGOAL (trace_goalno_tac assume_tac));
 | 
|  |    502 | 
 | 
|  |    503 | 
 | 
| 19053 |    504 | (** theorem database **)
 | 
| 18120 |    505 | 
 | 
|  |    506 | fun bind_thm (name, thm) = ThmDatabase.ml_store_thm (name, standard thm);
 | 
|  |    507 | fun bind_thms (name, thms) = ThmDatabase.ml_store_thms (name, map standard thms);
 | 
|  |    508 | 
 | 
|  |    509 | fun qed name = ThmDatabase.ml_store_thm (name, result ());
 | 
|  |    510 | fun qed_goal name thy goal tacsf = ThmDatabase.ml_store_thm (name, prove_goal thy goal tacsf);
 | 
|  |    511 | fun qed_goalw name thy rews goal tacsf =
 | 
|  |    512 |   ThmDatabase.ml_store_thm (name, prove_goalw thy rews goal tacsf);
 | 
|  |    513 | fun qed_spec_mp name =
 | 
|  |    514 |   ThmDatabase.ml_store_thm (name, ObjectLogic.rulify_no_asm (result ()));
 | 
|  |    515 | fun qed_goal_spec_mp name thy s p =
 | 
|  |    516 |   bind_thm (name, ObjectLogic.rulify_no_asm (prove_goal thy s p));
 | 
|  |    517 | fun qed_goalw_spec_mp name thy defs s p =
 | 
|  |    518 |   bind_thm (name, ObjectLogic.rulify_no_asm (prove_goalw thy defs s p));
 | 
|  |    519 | 
 | 
|  |    520 | fun no_qed () = ();
 | 
|  |    521 | 
 | 
|  |    522 | end;
 | 
|  |    523 | 
 | 
|  |    524 | structure Goals: GOALS = OldGoals;
 | 
|  |    525 | open Goals;
 |