wenzelm@18120: (* Title: Pure/old_goals.ML wenzelm@18120: ID: $Id$ wenzelm@18120: Author: Lawrence C Paulson, Cambridge University Computer Laboratory wenzelm@18120: Copyright 1993 University of Cambridge wenzelm@18120: wenzelm@18120: Old-style goal stack package. The goal stack initially holds a dummy wenzelm@18120: proof, and can never become empty. Each goal stack consists of a list wenzelm@18120: of levels. The undo list is a list of goal stacks. Finally, there wenzelm@18120: may be a stack of pending proofs. wenzelm@18120: *) wenzelm@18120: wenzelm@18120: signature GOALS = wenzelm@18120: sig wenzelm@26429: val the_context: unit -> theory wenzelm@18120: val premises: unit -> thm list wenzelm@18120: val prove_goal: theory -> string -> (thm list -> tactic list) -> thm wenzelm@18120: val prove_goalw: theory -> thm list -> string -> (thm list -> tactic list) -> thm wenzelm@18120: val topthm: unit -> thm wenzelm@18120: val result: unit -> thm wenzelm@18120: val uresult: unit -> thm wenzelm@18120: val getgoal: int -> term wenzelm@18120: val gethyps: int -> thm list wenzelm@18120: val prlev: int -> unit wenzelm@18120: val pr: unit -> unit wenzelm@18120: val prlim: int -> unit wenzelm@18120: val goal: theory -> string -> thm list wenzelm@18120: val goalw: theory -> thm list -> string -> thm list wenzelm@18120: val Goal: string -> thm list wenzelm@18120: val Goalw: thm list -> string -> thm list wenzelm@18120: val by: tactic -> unit wenzelm@18120: val back: unit -> unit wenzelm@18120: val choplev: int -> unit wenzelm@18120: val undo: unit -> unit wenzelm@18120: val qed: string -> unit wenzelm@18120: val qed_goal: string -> theory -> string -> (thm list -> tactic list) -> unit wenzelm@18120: val qed_goalw: string -> theory -> thm list -> string wenzelm@18120: -> (thm list -> tactic list) -> unit wenzelm@18120: val qed_spec_mp: string -> unit wenzelm@18120: val qed_goal_spec_mp: string -> theory -> string -> (thm list -> tactic list) -> unit wenzelm@18120: val qed_goalw_spec_mp: string -> theory -> thm list -> string wenzelm@18120: -> (thm list -> tactic list) -> unit wenzelm@18120: end; wenzelm@18120: wenzelm@18120: signature OLD_GOALS = wenzelm@18120: sig wenzelm@18120: include GOALS wenzelm@27256: val simple_read_term: theory -> typ -> string -> term wenzelm@27256: val read_term: theory -> string -> term wenzelm@27256: val read_prop: theory -> string -> term wenzelm@18120: type proof wenzelm@19053: val chop: unit -> unit wenzelm@18120: val reset_goals: unit -> unit wenzelm@18120: val result_error_fn: (thm -> string -> thm) ref wenzelm@18120: val print_sign_exn: theory -> exn -> 'a wenzelm@18120: val prove_goalw_cterm: thm list->cterm->(thm list->tactic list)->thm wenzelm@18120: val prove_goalw_cterm_nocheck: thm list->cterm->(thm list->tactic list)->thm wenzelm@18120: val print_exn: exn -> 'a wenzelm@18120: val filter_goal: (term*term->bool) -> thm list -> int -> thm list wenzelm@18120: val goalw_cterm: thm list -> cterm -> thm list wenzelm@18120: val simple_prove_goal_cterm: cterm->(thm list->tactic list)->thm wenzelm@18120: val byev: tactic list -> unit wenzelm@18120: val save_proof: unit -> proof wenzelm@18120: val restore_proof: proof -> thm list wenzelm@18120: val push_proof: unit -> unit wenzelm@18120: val pop_proof: unit -> thm list wenzelm@18120: val rotate_proof: unit -> thm list wenzelm@18120: end; wenzelm@18120: wenzelm@18120: structure OldGoals: OLD_GOALS = wenzelm@18120: struct wenzelm@18120: wenzelm@27256: (* global context *) wenzelm@27256: wenzelm@26429: val the_context = ML_Context.the_global_context; wenzelm@26429: wenzelm@26429: wenzelm@27256: (* old ways of reading terms *) wenzelm@27256: wenzelm@27256: fun simple_read_term thy T s = wenzelm@27256: let wenzelm@27256: val ctxt = ProofContext.init thy wenzelm@27256: |> ProofContext.allow_dummies wenzelm@27256: |> ProofContext.set_mode ProofContext.mode_schematic; wenzelm@27256: val parse = if T = propT then Syntax.parse_prop else Syntax.parse_term; wenzelm@27256: in parse ctxt s |> TypeInfer.constrain T |> Syntax.check_term ctxt end; wenzelm@27256: wenzelm@27256: fun read_term thy = simple_read_term thy dummyT; wenzelm@27256: fun read_prop thy = simple_read_term thy propT; wenzelm@27256: wenzelm@27256: wenzelm@27256: wenzelm@18120: (*** Goal package ***) wenzelm@18120: wenzelm@18120: (*Each level of goal stack includes a proof state and alternative states, wenzelm@18120: the output of the tactic applied to the preceeding level. *) wenzelm@18120: type gstack = (thm * thm Seq.seq) list; wenzelm@18120: wenzelm@18120: datatype proof = Proof of gstack list * thm list * (bool*thm->thm); wenzelm@18120: wenzelm@18120: wenzelm@18120: (*** References ***) wenzelm@18120: wenzelm@18120: (*Current assumption list -- set by "goal".*) wenzelm@18120: val curr_prems = ref([] : thm list); wenzelm@18120: wenzelm@18120: (*Return assumption list -- useful if you didn't save "goal"'s result. *) wenzelm@18120: fun premises() = !curr_prems; wenzelm@18120: wenzelm@18120: (*Current result maker -- set by "goal", used by "result". *) wenzelm@18120: fun init_mkresult _ = error "No goal has been supplied in subgoal module"; wenzelm@18120: val curr_mkresult = ref (init_mkresult: bool*thm->thm); wenzelm@18120: wenzelm@18120: (*List of previous goal stacks, for the undo operation. Set by setstate. wenzelm@18120: A list of lists!*) wenzelm@26429: val undo_list = ref([[(asm_rl, Seq.empty)]] : gstack list); wenzelm@18120: wenzelm@18120: (* Stack of proof attempts *) wenzelm@18120: val proofstack = ref([]: proof list); wenzelm@18120: wenzelm@18120: (*reset all refs*) wenzelm@18120: fun reset_goals () = wenzelm@18120: (curr_prems := []; curr_mkresult := init_mkresult; wenzelm@26429: undo_list := [[(asm_rl, Seq.empty)]]); wenzelm@18120: wenzelm@18120: wenzelm@18120: (*** Setting up goal-directed proof ***) wenzelm@18120: wenzelm@18120: (*Generates the list of new theories when the proof state's theory changes*) wenzelm@18120: fun thy_error (thy,thy') = wenzelm@29091: let val names = Context.display_names thy' \\ Context.display_names thy wenzelm@18120: in case names of wenzelm@18120: [name] => "\nNew theory: " ^ name wenzelm@18120: | _ => "\nNew theories: " ^ space_implode ", " names wenzelm@18120: end; wenzelm@18120: wenzelm@18120: (*Default action is to print an error message; could be suppressed for wenzelm@18120: special applications.*) wenzelm@18120: fun result_error_default state msg : thm = wenzelm@18120: Pretty.str "Bad final proof state:" :: Display.pretty_goals (!goals_limit) state @ wenzelm@18120: [Pretty.str msg, Pretty.str "Proof failed!"] |> Pretty.chunks |> Pretty.string_of |> error; wenzelm@18120: wenzelm@18120: val result_error_fn = ref result_error_default; wenzelm@18120: wenzelm@18120: wenzelm@18120: (*Common treatment of "goal" and "prove_goal": wenzelm@18120: Return assumptions, initial proof state, and function to make result. wenzelm@18120: "atomic" indicates if the goal should be wrapped up in the function wenzelm@18120: "Goal::prop=>prop" to avoid assumptions being returned separately. wenzelm@18120: *) wenzelm@18120: fun prepare_proof atomic rths chorn = wenzelm@18120: let wenzelm@26291: val _ = legacy_feature "Old goal command"; wenzelm@26626: val thy = Thm.theory_of_cterm chorn; wenzelm@26626: val horn = Thm.term_of chorn; wenzelm@18120: val _ = Term.no_dummy_patterns horn handle TERM (msg, _) => error msg; wenzelm@18120: val (As, B) = Logic.strip_horn horn; wenzelm@18120: val atoms = atomic andalso wenzelm@27332: forall (fn t => not (can Logic.dest_implies t orelse Logic.is_all t)) As; wenzelm@18120: val (As,B) = if atoms then ([],horn) else (As,B); wenzelm@18120: val cAs = map (cterm_of thy) As; wenzelm@26653: val prems = map (rewrite_rule rths o Thm.forall_elim_vars 0 o Thm.assume) cAs; wenzelm@18120: val cB = cterm_of thy B; wenzelm@18120: val st0 = let val st = Goal.init cB |> fold Thm.weaken cAs wenzelm@18120: in rewrite_goals_rule rths st end wenzelm@18120: (*discharges assumptions from state in the order they appear in goal; wenzelm@18120: checks (if requested) that resulting theorem is equivalent to goal. *) wenzelm@18120: fun mkresult (check,state) = wenzelm@18120: let val state = Seq.hd (flexflex_rule state) wenzelm@18120: handle THM _ => state (*smash flexflex pairs*) wenzelm@18120: val ngoals = nprems_of state wenzelm@18120: val ath = implies_intr_list cAs state wenzelm@18120: val th = Goal.conclude ath wenzelm@26626: val thy' = Thm.theory_of_thm th wenzelm@26626: val {hyps, prop, ...} = Thm.rep_thm th wenzelm@18120: val final_th = standard th wenzelm@18120: in if not check then final_th wenzelm@26665: else if not (Theory.eq_thy(thy,thy')) then !result_error_fn state wenzelm@18120: ("Theory of proof state has changed!" ^ wenzelm@18120: thy_error (thy,thy')) wenzelm@18120: else if ngoals>0 then !result_error_fn state wenzelm@18120: (string_of_int ngoals ^ " unsolved goals!") wenzelm@18120: else if not (null hyps) then !result_error_fn state wenzelm@18120: ("Additional hypotheses:\n" ^ wenzelm@26939: cat_lines (map (Syntax.string_of_term_global thy) hyps)) wenzelm@18120: else if Pattern.matches thy wenzelm@18120: (Envir.beta_norm (term_of chorn), Envir.beta_norm prop) wenzelm@18120: then final_th wenzelm@18120: else !result_error_fn state "proved a different theorem" wenzelm@18120: end wenzelm@18120: in wenzelm@26665: if Theory.eq_thy(thy, Thm.theory_of_thm st0) wenzelm@18120: then (prems, st0, mkresult) wenzelm@18120: else error ("Definitions would change the proof state's theory" ^ wenzelm@18120: thy_error (thy, Thm.theory_of_thm st0)) wenzelm@18120: end wenzelm@18120: handle THM(s,_,_) => error("prepare_proof: exception THM was raised!\n" ^ s); wenzelm@18120: wenzelm@18120: (*Prints exceptions readably to users*) wenzelm@18120: fun print_sign_exn_unit thy e = wenzelm@18120: case e of wenzelm@18120: THM (msg,i,thms) => wenzelm@18120: (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); wenzelm@26928: List.app Display.print_thm thms) wenzelm@18120: | THEORY (msg,thys) => wenzelm@18120: (writeln ("Exception THEORY raised:\n" ^ msg); wenzelm@18120: List.app (writeln o Context.str_of_thy) thys) wenzelm@18120: | TERM (msg,ts) => wenzelm@18120: (writeln ("Exception TERM raised:\n" ^ msg); wenzelm@26939: List.app (writeln o Syntax.string_of_term_global thy) ts) wenzelm@18120: | TYPE (msg,Ts,ts) => wenzelm@18120: (writeln ("Exception TYPE raised:\n" ^ msg); wenzelm@26939: List.app (writeln o Syntax.string_of_typ_global thy) Ts; wenzelm@26939: List.app (writeln o Syntax.string_of_term_global thy) ts) wenzelm@18120: | e => raise e; wenzelm@18120: wenzelm@18120: (*Prints an exception, then fails*) wenzelm@18678: fun print_sign_exn thy e = (print_sign_exn_unit thy e; raise ERROR ""); wenzelm@18120: wenzelm@18120: (** the prove_goal.... commands wenzelm@18120: Prove theorem using the listed tactics; check it has the specified form. wenzelm@18120: Augment theory with all type assignments of goal. wenzelm@18120: Syntax is similar to "goal" command for easy keyboard use. **) wenzelm@18120: wenzelm@18120: (*Version taking the goal as a cterm*) wenzelm@18120: fun prove_goalw_cterm_general check rths chorn tacsf = wenzelm@18120: let val (prems, st0, mkresult) = prepare_proof false rths chorn wenzelm@18120: val tac = EVERY (tacsf prems) wenzelm@18120: fun statef() = wenzelm@18120: (case Seq.pull (tac st0) of wenzelm@18120: SOME(st,_) => st wenzelm@18120: | _ => error ("prove_goal: tactic failed")) wenzelm@25685: in mkresult (check, cond_timeit (!Output.timing) "" statef) end wenzelm@26626: handle e => (print_sign_exn_unit (Thm.theory_of_cterm chorn) e; wenzelm@18120: writeln ("The exception above was raised for\n" ^ wenzelm@18120: Display.string_of_cterm chorn); raise e); wenzelm@18120: wenzelm@18120: (*Two variants: one checking the result, one not. wenzelm@18120: Neither prints runtime messages: they are for internal packages.*) wenzelm@18120: fun prove_goalw_cterm rths chorn = wenzelm@18120: setmp Output.timing false (prove_goalw_cterm_general true rths chorn) wenzelm@18120: and prove_goalw_cterm_nocheck rths chorn = wenzelm@18120: setmp Output.timing false (prove_goalw_cterm_general false rths chorn); wenzelm@18120: wenzelm@18120: wenzelm@18120: (*Version taking the goal as a string*) wenzelm@18120: fun prove_goalw thy rths agoal tacsf = wenzelm@27256: let val chorn = cterm_of thy (read_prop thy agoal) wenzelm@18120: in prove_goalw_cterm_general true rths chorn tacsf end wenzelm@27256: handle ERROR msg => cat_error msg (*from read_prop?*) wenzelm@18120: ("The error(s) above occurred for " ^ quote agoal); wenzelm@18120: wenzelm@18120: (*String version with no meta-rewrite-rules*) wenzelm@18120: fun prove_goal thy = prove_goalw thy []; wenzelm@18120: wenzelm@18120: wenzelm@18120: wenzelm@18120: (*** Commands etc ***) wenzelm@18120: wenzelm@18120: (*Return the current goal stack, if any, from undo_list*) wenzelm@18120: fun getstate() : gstack = case !undo_list of wenzelm@18120: [] => error"No current state in subgoal module" wenzelm@18120: | x::_ => x; wenzelm@18120: wenzelm@18120: (*Pops the given goal stack*) wenzelm@18120: fun pop [] = error"Cannot go back past the beginning of the proof!" wenzelm@18120: | pop (pair::pairs) = (pair,pairs); wenzelm@18120: wenzelm@18120: wenzelm@23635: (* Print a level of the goal stack *) wenzelm@18120: wenzelm@18120: fun print_top ((th, _), pairs) = wenzelm@23635: let wenzelm@23635: val n = length pairs; wenzelm@23635: val m = (! goals_limit); wenzelm@23635: val ngoals = nprems_of th; wenzelm@23635: in wenzelm@23635: [Pretty.str ("Level " ^ string_of_int n ^ wenzelm@23635: (if ngoals > 0 then " (" ^ string_of_int ngoals ^ " subgoal" ^ wenzelm@23635: (if ngoals <> 1 then "s" else "") ^ ")" wenzelm@23635: else ""))] @ wenzelm@23635: Display.pretty_goals m th wenzelm@23635: end |> Pretty.chunks |> Pretty.writeln; wenzelm@18120: wenzelm@18120: (*Printing can raise exceptions, so the assignment occurs last. wenzelm@18120: Can do setstate[(st,Seq.empty)] to set st as the state. *) wenzelm@18120: fun setstate newgoals = wenzelm@18120: (print_top (pop newgoals); undo_list := newgoals :: !undo_list); wenzelm@18120: wenzelm@18120: (*Given a proof state transformation, return a command that updates wenzelm@18120: the goal stack*) wenzelm@18120: fun make_command com = setstate (com (pop (getstate()))); wenzelm@18120: wenzelm@18120: (*Apply a function on proof states to the current goal stack*) wenzelm@18120: fun apply_fun f = f (pop(getstate())); wenzelm@18120: wenzelm@18120: (*Return the top theorem, representing the proof state*) wenzelm@18120: fun topthm () = apply_fun (fn ((th,_), _) => th); wenzelm@18120: wenzelm@18120: (*Return the final result. *) wenzelm@18120: fun result () = !curr_mkresult (true, topthm()); wenzelm@18120: wenzelm@18120: (*Return the result UNCHECKED that it equals the goal -- for synthesis, wenzelm@18120: answer extraction, or other instantiation of Vars *) wenzelm@18120: fun uresult () = !curr_mkresult (false, topthm()); wenzelm@18120: wenzelm@18120: (*Get subgoal i from goal stack*) wenzelm@18120: fun getgoal i = Logic.get_goal (prop_of (topthm())) i; wenzelm@18120: wenzelm@18120: (*Return subgoal i's hypotheses as meta-level assumptions. wenzelm@18120: For debugging uses of METAHYPS*) wenzelm@18120: local exception GETHYPS of thm list wenzelm@18120: in wenzelm@18120: fun gethyps i = wenzelm@18120: (METAHYPS (fn hyps => raise (GETHYPS hyps)) i (topthm()); []) wenzelm@18120: handle GETHYPS hyps => hyps wenzelm@18120: end; wenzelm@18120: wenzelm@18120: (*Prints exceptions nicely at top level; wenzelm@18120: raises the exception in order to have a polymorphic type!*) wenzelm@18120: fun print_exn e = (print_sign_exn_unit (Thm.theory_of_thm (topthm())) e; raise e); wenzelm@18120: wenzelm@18120: (*Which thms could apply to goal i? (debugs tactics involving filter_thms) *) wenzelm@18120: fun filter_goal could ths i = filter_thms could (999, getgoal i, ths); wenzelm@18120: wenzelm@18120: (*For inspecting earlier levels of the backward proof*) wenzelm@18120: fun chop_level n (pair,pairs) = wenzelm@18120: let val level = length pairs wenzelm@18120: in if n<0 andalso ~n <= level wenzelm@18120: then List.drop (pair::pairs, ~n) wenzelm@18120: else if 0<=n andalso n<= level wenzelm@18120: then List.drop (pair::pairs, level - n) wenzelm@18120: else error ("Level number must lie between 0 and " ^ wenzelm@18120: string_of_int level) wenzelm@18120: end; wenzelm@18120: wenzelm@18120: (*Print the given level of the proof; prlev ~1 prints previous level*) wenzelm@23635: fun prlev n = apply_fun (print_top o pop o (chop_level n)); wenzelm@23635: fun pr () = apply_fun print_top; wenzelm@18120: wenzelm@18120: (*Set goals_limit and print again*) wenzelm@18120: fun prlim n = (goals_limit:=n; pr()); wenzelm@18120: wenzelm@18120: (** the goal.... commands wenzelm@18120: Read main goal. Set global variables curr_prems, curr_mkresult. wenzelm@18120: Initial subgoal and premises are rewritten using rths. **) wenzelm@18120: wenzelm@18120: (*Version taking the goal as a cterm; if you have a term t and theory thy, use wenzelm@18120: goalw_cterm rths (cterm_of thy t); *) wenzelm@18120: fun agoalw_cterm atomic rths chorn = wenzelm@18120: let val (prems, st0, mkresult) = prepare_proof atomic rths chorn wenzelm@18120: in undo_list := []; wenzelm@18120: setstate [ (st0, Seq.empty) ]; wenzelm@18120: curr_prems := prems; wenzelm@18120: curr_mkresult := mkresult; wenzelm@18120: prems wenzelm@18120: end; wenzelm@18120: wenzelm@18120: val goalw_cterm = agoalw_cterm false; wenzelm@18120: wenzelm@18120: (*Version taking the goal as a string*) wenzelm@18120: fun agoalw atomic thy rths agoal = wenzelm@27256: agoalw_cterm atomic rths (cterm_of thy (read_prop thy agoal)) wenzelm@18678: handle ERROR msg => cat_error msg (*from type_assign, etc via prepare_proof*) wenzelm@18120: ("The error(s) above occurred for " ^ quote agoal); wenzelm@18120: wenzelm@18120: val goalw = agoalw false; wenzelm@18120: fun goal thy = goalw thy []; wenzelm@18120: wenzelm@18120: (*now the versions that wrap the goal up in `Goal' to make it atomic*) wenzelm@26429: fun Goalw thms s = agoalw true (ML_Context.the_global_context ()) thms s; wenzelm@18120: val Goal = Goalw []; wenzelm@18120: wenzelm@18120: (*simple version with minimal amount of checking and postprocessing*) wenzelm@18120: fun simple_prove_goal_cterm G f = wenzelm@18120: let wenzelm@26291: val _ = legacy_feature "Old goal command"; wenzelm@18120: val As = Drule.strip_imp_prems G; wenzelm@18120: val B = Drule.strip_imp_concl G; wenzelm@20229: val asms = map Assumption.assume As; wenzelm@18120: fun check NONE = error "prove_goal: tactic failed" wenzelm@18120: | check (SOME (thm, _)) = (case nprems_of thm of wenzelm@18120: 0 => thm wenzelm@18120: | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!")) wenzelm@18120: in wenzelm@18120: standard (implies_intr_list As wenzelm@18120: (check (Seq.pull (EVERY (f asms) (trivial B))))) wenzelm@18120: end; wenzelm@18120: wenzelm@18120: wenzelm@18120: (*Proof step "by" the given tactic -- apply tactic to the proof state*) wenzelm@18120: fun by_com tac ((th,ths), pairs) : gstack = wenzelm@18120: (case Seq.pull(tac th) of wenzelm@18120: NONE => error"by: tactic failed" wenzelm@18120: | SOME(th2,ths2) => wenzelm@22360: (if Thm.eq_thm(th,th2) wenzelm@18120: then warning "Warning: same as previous level" wenzelm@22360: else if Thm.eq_thm_thy(th,th2) then () wenzelm@18120: else warning ("Warning: theory of proof state has changed" ^ wenzelm@18120: thy_error (Thm.theory_of_thm th, Thm.theory_of_thm th2)); wenzelm@18120: ((th2,ths2)::(th,ths)::pairs))); wenzelm@18120: wenzelm@25685: fun by tac = cond_timeit (!Output.timing) "" wenzelm@18120: (fn() => make_command (by_com tac)); wenzelm@18120: wenzelm@18120: (* byev[tac1,...,tacn] applies tac1 THEN ... THEN tacn. wenzelm@18120: Good for debugging proofs involving prove_goal.*) wenzelm@18120: val byev = by o EVERY; wenzelm@18120: wenzelm@18120: wenzelm@18120: (*Backtracking means find an alternative result from a tactic. wenzelm@18120: If none at this level, try earlier levels*) wenzelm@18120: fun backtrack [] = error"back: no alternatives" wenzelm@18120: | backtrack ((th,thstr) :: pairs) = wenzelm@18120: (case Seq.pull thstr of wenzelm@18120: NONE => (writeln"Going back a level..."; backtrack pairs) wenzelm@18120: | SOME(th2,thstr2) => wenzelm@22360: (if Thm.eq_thm(th,th2) wenzelm@18120: then warning "Warning: same as previous choice at this level" wenzelm@22360: else if Thm.eq_thm_thy(th,th2) then () wenzelm@18120: else warning "Warning: theory of proof state has changed"; wenzelm@18120: (th2,thstr2)::pairs)); wenzelm@18120: wenzelm@18120: fun back() = setstate (backtrack (getstate())); wenzelm@18120: wenzelm@18120: (*Chop back to previous level of the proof*) wenzelm@18120: fun choplev n = make_command (chop_level n); wenzelm@18120: wenzelm@18120: (*Chopping back the goal stack*) wenzelm@18120: fun chop () = make_command (fn (_,pairs) => pairs); wenzelm@18120: wenzelm@18120: (*Restore the previous proof state; discard current state. *) wenzelm@18120: fun undo() = case !undo_list of wenzelm@18120: [] => error"No proof state" wenzelm@18120: | [_] => error"Already at initial state" wenzelm@18120: | _::newundo => (undo_list := newundo; pr()) ; wenzelm@18120: wenzelm@18120: wenzelm@18120: (*** Managing the proof stack ***) wenzelm@18120: wenzelm@18120: fun save_proof() = Proof(!undo_list, !curr_prems, !curr_mkresult); wenzelm@18120: wenzelm@18120: fun restore_proof(Proof(ul,prems,mk)) = wenzelm@18120: (undo_list:= ul; curr_prems:= prems; curr_mkresult := mk; prems); wenzelm@18120: wenzelm@18120: wenzelm@18120: fun top_proof() = case !proofstack of wenzelm@18120: [] => error("Stack of proof attempts is empty!") wenzelm@18120: | p::ps => (p,ps); wenzelm@18120: wenzelm@18120: (* push a copy of the current proof state on to the stack *) wenzelm@18120: fun push_proof() = (proofstack := (save_proof() :: !proofstack)); wenzelm@18120: wenzelm@18120: (* discard the top proof state of the stack *) wenzelm@18120: fun pop_proof() = wenzelm@18120: let val (p,ps) = top_proof() wenzelm@18120: val prems = restore_proof p wenzelm@18120: in proofstack := ps; pr(); prems end; wenzelm@18120: wenzelm@18120: (* rotate the stack so that the top element goes to the bottom *) wenzelm@18120: fun rotate_proof() = let val (p,ps) = top_proof() wenzelm@18120: in proofstack := ps@[save_proof()]; wenzelm@18120: restore_proof p; wenzelm@18120: pr(); wenzelm@18120: !curr_prems wenzelm@18120: end; wenzelm@18120: wenzelm@18120: wenzelm@26414: (** theorem bindings **) wenzelm@18120: wenzelm@26414: fun qed name = ML_Context.ml_store_thm (name, result ()); wenzelm@26414: fun qed_goal name thy goal tacsf = ML_Context.ml_store_thm (name, prove_goal thy goal tacsf); wenzelm@18120: fun qed_goalw name thy rews goal tacsf = wenzelm@26414: ML_Context.ml_store_thm (name, prove_goalw thy rews goal tacsf); wenzelm@18120: fun qed_spec_mp name = wenzelm@26414: ML_Context.ml_store_thm (name, ObjectLogic.rulify_no_asm (result ())); wenzelm@18120: fun qed_goal_spec_mp name thy s p = wenzelm@18120: bind_thm (name, ObjectLogic.rulify_no_asm (prove_goal thy s p)); wenzelm@18120: fun qed_goalw_spec_mp name thy defs s p = wenzelm@18120: bind_thm (name, ObjectLogic.rulify_no_asm (prove_goalw thy defs s p)); wenzelm@18120: wenzelm@18120: end; wenzelm@18120: wenzelm@18120: structure Goals: GOALS = OldGoals; wenzelm@18120: open Goals;