author | wenzelm |
Mon, 22 Oct 2001 18:01:38 +0200 | |
changeset 11886 | 36d0585f87de |
parent 11884 | 341b1fbc6130 |
child 11963 | a6608d44a46b |
permissions | -rw-r--r-- |
1460 | 1 |
(* Title: Pure/goals.ML |
0 | 2 |
ID: $Id$ |
1460 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
6 |
Goal stack package. The goal stack initially holds a dummy proof, and can |
|
7 |
never become empty. Each goal stack consists of a list of levels. The |
|
8 |
undo list is a list of goal stacks. Finally, there may be a stack of |
|
9 |
pending proofs. |
|
10 |
*) |
|
11 |
||
12 |
signature GOALS = |
|
11884 | 13 |
sig |
0 | 14 |
type proof |
7234 | 15 |
val reset_goals : unit -> unit |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
16 |
val atomic_goal : theory -> string -> thm list |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
17 |
val atomic_goalw : theory -> thm list -> string -> thm list |
6189 | 18 |
val Goal : string -> thm list |
19 |
val Goalw : thm list -> string -> thm list |
|
1460 | 20 |
val ba : int -> unit |
21 |
val back : unit -> unit |
|
22 |
val bd : thm -> int -> unit |
|
23 |
val bds : thm list -> int -> unit |
|
24 |
val be : thm -> int -> unit |
|
25 |
val bes : thm list -> int -> unit |
|
26 |
val br : thm -> int -> unit |
|
27 |
val brs : thm list -> int -> unit |
|
28 |
val bw : thm -> unit |
|
29 |
val bws : thm list -> unit |
|
30 |
val by : tactic -> unit |
|
31 |
val byev : tactic list -> unit |
|
32 |
val chop : unit -> unit |
|
33 |
val choplev : int -> unit |
|
6017 | 34 |
val export_thy : theory -> thm -> thm |
5246 | 35 |
val export : thm -> thm |
6189 | 36 |
val Export : thm -> thm |
1460 | 37 |
val fa : unit -> unit |
38 |
val fd : thm -> unit |
|
39 |
val fds : thm list -> unit |
|
40 |
val fe : thm -> unit |
|
41 |
val fes : thm list -> unit |
|
42 |
val filter_goal : (term*term->bool) -> thm list -> int -> thm list |
|
43 |
val fr : thm -> unit |
|
44 |
val frs : thm list -> unit |
|
45 |
val getgoal : int -> term |
|
46 |
val gethyps : int -> thm list |
|
47 |
val goal : theory -> string -> thm list |
|
48 |
val goalw : theory -> thm list -> string -> thm list |
|
49 |
val goalw_cterm : thm list -> cterm -> thm list |
|
50 |
val pop_proof : unit -> thm list |
|
51 |
val pr : unit -> unit |
|
8884 | 52 |
val disable_pr : unit -> unit |
53 |
val enable_pr : unit -> unit |
|
2580
e3f680709487
Gradual switching to Basis Library functions nth, drop, etc.
paulson
parents:
2514
diff
changeset
|
54 |
val prlev : int -> unit |
2514 | 55 |
val prlim : int -> unit |
1460 | 56 |
val premises : unit -> thm list |
57 |
val prin : term -> unit |
|
58 |
val printyp : typ -> unit |
|
59 |
val pprint_term : term -> pprint_args -> unit |
|
60 |
val pprint_typ : typ -> pprint_args -> unit |
|
61 |
val print_exn : exn -> 'a |
|
62 |
val print_sign_exn : Sign.sg -> exn -> 'a |
|
63 |
val prove_goal : theory -> string -> (thm list -> tactic list) -> thm |
|
577 | 64 |
val prove_goalw : theory->thm list->string->(thm list->tactic list)->thm |
1460 | 65 |
val prove_goalw_cterm : thm list->cterm->(thm list->tactic list)->thm |
5311
f3f71669878e
Rule mk_triv_goal for making instances of triv_goal
paulson
parents:
5246
diff
changeset
|
66 |
val prove_goalw_cterm_nocheck : thm list->cterm->(thm list->tactic list)->thm |
11884 | 67 |
val quick_and_dirty_prove_goalw_cterm: theory -> thm list -> cterm |
68 |
-> (thm list -> tactic list) -> thm |
|
1460 | 69 |
val push_proof : unit -> unit |
70 |
val read : string -> term |
|
71 |
val ren : string -> int -> unit |
|
72 |
val restore_proof : proof -> thm list |
|
73 |
val result : unit -> thm |
|
3532 | 74 |
val result_error_fn : (thm -> string -> thm) ref |
1460 | 75 |
val rotate_proof : unit -> thm list |
76 |
val uresult : unit -> thm |
|
77 |
val save_proof : unit -> proof |
|
78 |
val topthm : unit -> thm |
|
79 |
val undo : unit -> unit |
|
11884 | 80 |
val bind_thm : string * thm -> unit |
81 |
val bind_thms : string * thm list -> unit |
|
82 |
val qed : string -> unit |
|
83 |
val qed_goal : string -> theory -> string -> (thm list -> tactic list) -> unit |
|
84 |
val qed_goalw : string -> theory -> thm list -> string |
|
85 |
-> (thm list -> tactic list) -> unit |
|
86 |
val qed_spec_mp : string -> unit |
|
87 |
val qed_goal_spec_mp : string -> theory -> string -> (thm list -> tactic list) -> unit |
|
88 |
val qed_goalw_spec_mp : string -> theory -> thm list -> string |
|
89 |
-> (thm list -> tactic list) -> unit |
|
90 |
val no_qed: unit -> unit |
|
91 |
val thms_containing : xstring list -> (string * thm) list |
|
92 |
val findI : int -> (string * thm) list |
|
93 |
val findEs : int -> (string * thm) list |
|
94 |
val findE : int -> int -> (string * thm) list |
|
95 |
end; |
|
0 | 96 |
|
1500 | 97 |
structure Goals : GOALS = |
0 | 98 |
struct |
99 |
||
100 |
(*Each level of goal stack includes a proof state and alternative states, |
|
101 |
the output of the tactic applied to the preceeding level. *) |
|
4270 | 102 |
type gstack = (thm * thm Seq.seq) list; |
0 | 103 |
|
104 |
datatype proof = Proof of gstack list * thm list * (bool*thm->thm); |
|
105 |
||
5246 | 106 |
|
0 | 107 |
(*** References ***) |
108 |
||
109 |
(*Current assumption list -- set by "goal".*) |
|
110 |
val curr_prems = ref([] : thm list); |
|
111 |
||
112 |
(*Return assumption list -- useful if you didn't save "goal"'s result. *) |
|
113 |
fun premises() = !curr_prems; |
|
114 |
||
115 |
(*Current result maker -- set by "goal", used by "result". *) |
|
7234 | 116 |
fun init_mkresult _ = error "No goal has been supplied in subgoal module"; |
117 |
val curr_mkresult = ref (init_mkresult: bool*thm->thm); |
|
0 | 118 |
|
6390 | 119 |
val dummy = trivial(read_cterm (Theory.sign_of ProtoPure.thy) |
0 | 120 |
("PROP No_goal_has_been_supplied",propT)); |
121 |
||
122 |
(*List of previous goal stacks, for the undo operation. Set by setstate. |
|
123 |
A list of lists!*) |
|
4270 | 124 |
val undo_list = ref([[(dummy, Seq.empty)]] : gstack list); |
0 | 125 |
|
126 |
(* Stack of proof attempts *) |
|
127 |
val proofstack = ref([]: proof list); |
|
128 |
||
7234 | 129 |
(*reset all refs*) |
130 |
fun reset_goals () = |
|
131 |
(curr_prems := []; curr_mkresult := init_mkresult; |
|
7942 | 132 |
undo_list := [[(dummy, Seq.empty)]]); |
7234 | 133 |
|
0 | 134 |
|
135 |
(*** Setting up goal-directed proof ***) |
|
136 |
||
137 |
(*Generates the list of new theories when the proof state's signature changes*) |
|
138 |
fun sign_error (sign,sign') = |
|
3974 | 139 |
let val names = Sign.stamp_names_of sign' \\ Sign.stamp_names_of sign |
140 |
in case names of |
|
141 |
[name] => "\nNew theory: " ^ name |
|
142 |
| _ => "\nNew theories: " ^ space_implode ", " names |
|
0 | 143 |
end; |
144 |
||
1928 | 145 |
(*Default action is to print an error message; could be suppressed for |
146 |
special applications.*) |
|
3669
3384c6f1f095
removed print_goals_ref (which was broken anyway);
wenzelm
parents:
3536
diff
changeset
|
147 |
fun result_error_default state msg : thm = |
11884 | 148 |
Pretty.str "Bad final proof state:" :: Display.pretty_goals (!goals_limit) state @ |
11562 | 149 |
[Pretty.str msg, Pretty.str "Proof failed!"] |> Pretty.chunks |> Pretty.string_of |> error; |
1928 | 150 |
|
3532 | 151 |
val result_error_fn = ref result_error_default; |
1928 | 152 |
|
5246 | 153 |
(* alternative to standard: this function does not discharge the hypotheses |
154 |
first. Is used for locales, in the following function prepare_proof *) |
|
155 |
fun varify th = |
|
156 |
let val {maxidx,...} = rep_thm th |
|
157 |
in |
|
158 |
th |> forall_intr_frees |> forall_elim_vars (maxidx + 1) |
|
7637 | 159 |
|> Drule.strip_shyps_warning |
5246 | 160 |
|> zero_var_indexes |> Thm.varifyT |> Thm.compress |
161 |
end; |
|
162 |
||
163 |
(** exporting a theorem out of a locale means turning all meta-hyps into assumptions |
|
6017 | 164 |
and expand and cancel the locale definitions. |
165 |
export goes through all levels in case of nested locales, whereas |
|
6189 | 166 |
export_thy just goes one up. **) |
6017 | 167 |
|
168 |
fun get_top_scope_thms thy = |
|
6390 | 169 |
let val sc = (Locale.get_scope_sg (Theory.sign_of thy)) |
6017 | 170 |
in if (null sc) then (warning "No locale in theory"; []) |
171 |
else map (#prop o rep_thm) (map #2 (#thms(snd(hd sc)))) |
|
172 |
end; |
|
173 |
||
174 |
fun implies_intr_some_hyps thy A_set th = |
|
175 |
let |
|
6390 | 176 |
val sign = Theory.sign_of thy; |
6017 | 177 |
val used_As = A_set inter #hyps(rep_thm(th)); |
178 |
val ctl = map (cterm_of sign) used_As |
|
179 |
in foldl (fn (x, y) => Thm.implies_intr y x) (th, ctl) |
|
180 |
end; |
|
181 |
||
182 |
fun standard_some thy A_set th = |
|
183 |
let val {maxidx,...} = rep_thm th |
|
184 |
in |
|
185 |
th |> implies_intr_some_hyps thy A_set |
|
186 |
|> forall_intr_frees |> forall_elim_vars (maxidx + 1) |
|
7637 | 187 |
|> Drule.strip_shyps_warning |
6017 | 188 |
|> zero_var_indexes |> Thm.varifyT |> Thm.compress |
189 |
end; |
|
190 |
||
191 |
fun export_thy thy th = |
|
192 |
let val A_set = get_top_scope_thms thy |
|
193 |
in |
|
194 |
standard_some thy [] (Seq.hd ((REPEAT (FIRSTGOAL (rtac reflexive_thm))) (standard_some thy A_set th))) |
|
195 |
end; |
|
196 |
||
5246 | 197 |
val export = standard o Seq.hd o (REPEAT (FIRSTGOAL (rtac reflexive_thm))) o standard; |
198 |
||
6189 | 199 |
fun Export th = export_thy (the_context ()) th; |
200 |
||
201 |
||
0 | 202 |
(*Common treatment of "goal" and "prove_goal": |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
203 |
Return assumptions, initial proof state, and function to make result. |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
204 |
"atomic" indicates if the goal should be wrapped up in the function |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
205 |
"Goal::prop=>prop" to avoid assumptions being returned separately. |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
206 |
*) |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
207 |
fun prepare_proof atomic rths chorn = |
230 | 208 |
let val {sign, t=horn,...} = rep_cterm chorn; |
9533 | 209 |
val _ = Term.no_dummy_patterns horn handle TERM (msg, _) => error msg; |
0 | 210 |
val (_,As,B) = Logic.strip_horn(horn); |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
211 |
val atoms = atomic andalso |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
212 |
forall (fn t => not(Logic.is_implies t orelse Logic.is_all t)) As; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
213 |
val (As,B) = if atoms then ([],horn) else (As,B); |
230 | 214 |
val cAs = map (cterm_of sign) As; |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
215 |
val prems = map (rewrite_rule rths o forall_elim_vars(0) o assume) cAs; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
216 |
val cB = cterm_of sign B; |
5311
f3f71669878e
Rule mk_triv_goal for making instances of triv_goal
paulson
parents:
5246
diff
changeset
|
217 |
val st0 = let val st = Drule.mk_triv_goal cB |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
218 |
in rewrite_goals_rule rths st end |
0 | 219 |
(*discharges assumptions from state in the order they appear in goal; |
1460 | 220 |
checks (if requested) that resulting theorem is equivalent to goal. *) |
0 | 221 |
fun mkresult (check,state) = |
4270 | 222 |
let val state = Seq.hd (flexflex_rule state) |
1565 | 223 |
handle THM _ => state (*smash flexflex pairs*) |
1460 | 224 |
val ngoals = nprems_of state |
7637 | 225 |
val ath = implies_intr_list cAs state |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
226 |
val th = ath RS Drule.rev_triv_goal |
1240 | 227 |
val {hyps,prop,sign=sign',...} = rep_thm th |
9573
5cadc8146573
the "nocheck" versions of goal functions now standardize their result
paulson
parents:
9533
diff
changeset
|
228 |
val final_th = if (null(hyps)) then standard th else varify th |
5cadc8146573
the "nocheck" versions of goal functions now standardize their result
paulson
parents:
9533
diff
changeset
|
229 |
in if not check then final_th |
3532 | 230 |
else if not (Sign.eq_sg(sign,sign')) then !result_error_fn state |
1460 | 231 |
("Signature of proof state has changed!" ^ |
232 |
sign_error (sign,sign')) |
|
3532 | 233 |
else if ngoals>0 then !result_error_fn state |
1460 | 234 |
(string_of_int ngoals ^ " unsolved goals!") |
5246 | 235 |
else if (not (null hyps) andalso (not (Locale.in_locale hyps sign))) |
236 |
then !result_error_fn state |
|
5748
5a571d57183f
better reporting of "Additional hypotheses" in a locale
paulson
parents:
5729
diff
changeset
|
237 |
("Additional hypotheses:\n" ^ |
5a571d57183f
better reporting of "Additional hypotheses" in a locale
paulson
parents:
5729
diff
changeset
|
238 |
cat_lines |
5a571d57183f
better reporting of "Additional hypotheses" in a locale
paulson
parents:
5729
diff
changeset
|
239 |
(map (Sign.string_of_term sign) |
5a571d57183f
better reporting of "Additional hypotheses" in a locale
paulson
parents:
5729
diff
changeset
|
240 |
(filter (fn x => (not (Locale.in_locale [x] sign))) |
5a571d57183f
better reporting of "Additional hypotheses" in a locale
paulson
parents:
5729
diff
changeset
|
241 |
hyps))) |
1460 | 242 |
else if Pattern.matches (#tsig(Sign.rep_sg sign)) |
10487 | 243 |
(Envir.beta_norm (term_of chorn), Envir.beta_norm prop) |
9573
5cadc8146573
the "nocheck" versions of goal functions now standardize their result
paulson
parents:
9533
diff
changeset
|
244 |
then final_th |
3532 | 245 |
else !result_error_fn state "proved a different theorem" |
0 | 246 |
end |
678
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
577
diff
changeset
|
247 |
in |
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
577
diff
changeset
|
248 |
if Sign.eq_sg(sign, #sign(rep_thm st0)) |
0 | 249 |
then (prems, st0, mkresult) |
250 |
else error ("Definitions would change the proof state's signature" ^ |
|
1460 | 251 |
sign_error (sign, #sign(rep_thm st0))) |
0 | 252 |
end |
253 |
handle THM(s,_,_) => error("prepare_proof: exception THM was raised!\n" ^ s); |
|
254 |
||
255 |
(*Prints exceptions readably to users*) |
|
577 | 256 |
fun print_sign_exn_unit sign e = |
0 | 257 |
case e of |
258 |
THM (msg,i,thms) => |
|
1460 | 259 |
(writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); |
260 |
seq print_thm thms) |
|
0 | 261 |
| THEORY (msg,thys) => |
1460 | 262 |
(writeln ("Exception THEORY raised:\n" ^ msg); |
4994 | 263 |
seq (Pretty.writeln o Display.pretty_theory) thys) |
0 | 264 |
| TERM (msg,ts) => |
1460 | 265 |
(writeln ("Exception TERM raised:\n" ^ msg); |
266 |
seq (writeln o Sign.string_of_term sign) ts) |
|
0 | 267 |
| TYPE (msg,Ts,ts) => |
1460 | 268 |
(writeln ("Exception TYPE raised:\n" ^ msg); |
269 |
seq (writeln o Sign.string_of_typ sign) Ts; |
|
270 |
seq (writeln o Sign.string_of_term sign) ts) |
|
0 | 271 |
| e => raise e; |
272 |
||
577 | 273 |
(*Prints an exception, then fails*) |
274 |
fun print_sign_exn sign e = (print_sign_exn_unit sign e; raise ERROR); |
|
275 |
||
0 | 276 |
(** the prove_goal.... commands |
277 |
Prove theorem using the listed tactics; check it has the specified form. |
|
278 |
Augment signature with all type assignments of goal. |
|
279 |
Syntax is similar to "goal" command for easy keyboard use. **) |
|
280 |
||
281 |
(*Version taking the goal as a cterm*) |
|
5311
f3f71669878e
Rule mk_triv_goal for making instances of triv_goal
paulson
parents:
5246
diff
changeset
|
282 |
fun prove_goalw_cterm_general check rths chorn tacsf = |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
283 |
let val (prems, st0, mkresult) = prepare_proof false rths chorn |
0 | 284 |
val tac = EVERY (tacsf prems) |
285 |
fun statef() = |
|
4270 | 286 |
(case Seq.pull (tac st0) of |
1460 | 287 |
Some(st,_) => st |
288 |
| _ => error ("prove_goal: tactic failed")) |
|
8999 | 289 |
in mkresult (check, cond_timeit (!Library.timing) statef) end |
914
cae574c09137
Now prove_goalw_cterm calls print_sign_exn_unit, so that the rest
lcp
parents:
696
diff
changeset
|
290 |
handle e => (print_sign_exn_unit (#sign (rep_cterm chorn)) e; |
6960 | 291 |
writeln ("The exception above was raised for\n" ^ |
11884 | 292 |
Display.string_of_cterm chorn); raise e); |
545
fc4ff96bb0e9
Pure/goals.ML: prove_goalw_cterm now does its own exception-handling, moved
lcp
parents:
253
diff
changeset
|
293 |
|
5614
0e8b45a7d104
Now prove_goalw_cterm never prints timing statistics
paulson
parents:
5311
diff
changeset
|
294 |
(*Two variants: one checking the result, one not. |
0e8b45a7d104
Now prove_goalw_cterm never prints timing statistics
paulson
parents:
5311
diff
changeset
|
295 |
Neither prints runtime messages: they are for internal packages.*) |
0e8b45a7d104
Now prove_goalw_cterm never prints timing statistics
paulson
parents:
5311
diff
changeset
|
296 |
fun prove_goalw_cterm rths chorn = |
8999 | 297 |
setmp Library.timing false (prove_goalw_cterm_general true rths chorn) |
5614
0e8b45a7d104
Now prove_goalw_cterm never prints timing statistics
paulson
parents:
5311
diff
changeset
|
298 |
and prove_goalw_cterm_nocheck rths chorn = |
8999 | 299 |
setmp Library.timing false (prove_goalw_cterm_general false rths chorn); |
5311
f3f71669878e
Rule mk_triv_goal for making instances of triv_goal
paulson
parents:
5246
diff
changeset
|
300 |
|
0 | 301 |
|
302 |
(*Version taking the goal as a string*) |
|
303 |
fun prove_goalw thy rths agoal tacsf = |
|
6390 | 304 |
let val sign = Theory.sign_of thy |
230 | 305 |
val chorn = read_cterm sign (agoal,propT) |
5614
0e8b45a7d104
Now prove_goalw_cterm never prints timing statistics
paulson
parents:
5311
diff
changeset
|
306 |
in prove_goalw_cterm_general true rths chorn tacsf end |
545
fc4ff96bb0e9
Pure/goals.ML: prove_goalw_cterm now does its own exception-handling, moved
lcp
parents:
253
diff
changeset
|
307 |
handle ERROR => error (*from read_cterm?*) |
3536 | 308 |
("The error(s) above occurred for " ^ quote agoal); |
0 | 309 |
|
310 |
(*String version with no meta-rewrite-rules*) |
|
311 |
fun prove_goal thy = prove_goalw thy []; |
|
312 |
||
11884 | 313 |
(*quick and dirty version (conditional)*) |
314 |
fun quick_and_dirty_prove_goalw_cterm thy rews ct tacs = |
|
315 |
prove_goalw_cterm rews ct |
|
316 |
(if ! quick_and_dirty then (K [SkipProof.cheat_tac thy]) else tacs); |
|
317 |
||
0 | 318 |
|
319 |
(*** Commands etc ***) |
|
320 |
||
321 |
(*Return the current goal stack, if any, from undo_list*) |
|
322 |
fun getstate() : gstack = case !undo_list of |
|
323 |
[] => error"No current state in subgoal module" |
|
324 |
| x::_ => x; |
|
325 |
||
326 |
(*Pops the given goal stack*) |
|
327 |
fun pop [] = error"Cannot go back past the beginning of the proof!" |
|
328 |
| pop (pair::pairs) = (pair,pairs); |
|
329 |
||
330 |
||
8884 | 331 |
(* Print a level of the goal stack -- subject to quiet mode *) |
332 |
||
333 |
val quiet = ref false; |
|
334 |
fun disable_pr () = quiet := true; |
|
335 |
fun enable_pr () = quiet := false; |
|
336 |
||
3532 | 337 |
fun print_top ((th, _), pairs) = |
8884 | 338 |
if ! quiet then () |
11884 | 339 |
else ! Display.print_current_goals_fn (length pairs) (! goals_limit) th; |
0 | 340 |
|
341 |
(*Printing can raise exceptions, so the assignment occurs last. |
|
4270 | 342 |
Can do setstate[(st,Seq.empty)] to set st as the state. *) |
0 | 343 |
fun setstate newgoals = |
344 |
(print_top (pop newgoals); undo_list := newgoals :: !undo_list); |
|
345 |
||
346 |
(*Given a proof state transformation, return a command that updates |
|
347 |
the goal stack*) |
|
348 |
fun make_command com = setstate (com (pop (getstate()))); |
|
349 |
||
350 |
(*Apply a function on proof states to the current goal stack*) |
|
351 |
fun apply_fun f = f (pop(getstate())); |
|
352 |
||
353 |
(*Return the top theorem, representing the proof state*) |
|
354 |
fun topthm () = apply_fun (fn ((th,_), _) => th); |
|
355 |
||
356 |
(*Return the final result. *) |
|
357 |
fun result () = !curr_mkresult (true, topthm()); |
|
358 |
||
359 |
(*Return the result UNCHECKED that it equals the goal -- for synthesis, |
|
360 |
answer extraction, or other instantiation of Vars *) |
|
361 |
fun uresult () = !curr_mkresult (false, topthm()); |
|
362 |
||
363 |
(*Get subgoal i from goal stack*) |
|
2580
e3f680709487
Gradual switching to Basis Library functions nth, drop, etc.
paulson
parents:
2514
diff
changeset
|
364 |
fun getgoal i = List.nth (prems_of (topthm()), i-1) |
e3f680709487
Gradual switching to Basis Library functions nth, drop, etc.
paulson
parents:
2514
diff
changeset
|
365 |
handle Subscript => error"getgoal: Goal number out of range"; |
0 | 366 |
|
367 |
(*Return subgoal i's hypotheses as meta-level assumptions. |
|
368 |
For debugging uses of METAHYPS*) |
|
369 |
local exception GETHYPS of thm list |
|
370 |
in |
|
371 |
fun gethyps i = |
|
1500 | 372 |
(METAHYPS (fn hyps => raise (GETHYPS hyps)) i (topthm()); []) |
0 | 373 |
handle GETHYPS hyps => hyps |
374 |
end; |
|
375 |
||
376 |
(*Which thms could apply to goal i? (debugs tactics involving filter_thms) *) |
|
377 |
fun filter_goal could ths i = filter_thms could (999, getgoal i, ths); |
|
378 |
||
379 |
(*For inspecting earlier levels of the backward proof*) |
|
380 |
fun chop_level n (pair,pairs) = |
|
381 |
let val level = length pairs |
|
2126
d927beecedf8
Allowing negative levels (as offsets) in prlev and choplev
paulson
parents:
1928
diff
changeset
|
382 |
in if n<0 andalso ~n <= level |
2580
e3f680709487
Gradual switching to Basis Library functions nth, drop, etc.
paulson
parents:
2514
diff
changeset
|
383 |
then List.drop (pair::pairs, ~n) |
2126
d927beecedf8
Allowing negative levels (as offsets) in prlev and choplev
paulson
parents:
1928
diff
changeset
|
384 |
else if 0<=n andalso n<= level |
2580
e3f680709487
Gradual switching to Basis Library functions nth, drop, etc.
paulson
parents:
2514
diff
changeset
|
385 |
then List.drop (pair::pairs, level - n) |
0 | 386 |
else error ("Level number must lie between 0 and " ^ |
1460 | 387 |
string_of_int level) |
0 | 388 |
end; |
389 |
||
2514 | 390 |
(*Print the given level of the proof; prlev ~1 prints previous level*) |
8884 | 391 |
fun prlev n = (enable_pr (); apply_fun (print_top o pop o (chop_level n))); |
392 |
fun pr () = (enable_pr (); apply_fun print_top); |
|
0 | 393 |
|
2514 | 394 |
(*Set goals_limit and print again*) |
395 |
fun prlim n = (goals_limit:=n; pr()); |
|
396 |
||
0 | 397 |
(** the goal.... commands |
398 |
Read main goal. Set global variables curr_prems, curr_mkresult. |
|
399 |
Initial subgoal and premises are rewritten using rths. **) |
|
400 |
||
401 |
(*Version taking the goal as a cterm; if you have a term t and theory thy, use |
|
6390 | 402 |
goalw_cterm rths (cterm_of (Theory.sign_of thy) t); *) |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
403 |
fun agoalw_cterm atomic rths chorn = |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
404 |
let val (prems, st0, mkresult) = prepare_proof atomic rths chorn |
0 | 405 |
in undo_list := []; |
4270 | 406 |
setstate [ (st0, Seq.empty) ]; |
0 | 407 |
curr_prems := prems; |
408 |
curr_mkresult := mkresult; |
|
409 |
prems |
|
410 |
end; |
|
411 |
||
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
412 |
val goalw_cterm = agoalw_cterm false; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
413 |
|
0 | 414 |
(*Version taking the goal as a string*) |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
415 |
fun agoalw atomic thy rths agoal = |
6390 | 416 |
agoalw_cterm atomic rths (Locale.read_cterm(Theory.sign_of thy)(agoal,propT)) |
5246 | 417 |
handle ERROR => error (*from type_assign, etc via prepare_proof*) |
418 |
("The error(s) above occurred for " ^ quote agoal); |
|
0 | 419 |
|
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
420 |
val goalw = agoalw false; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
421 |
|
0 | 422 |
(*String version with no meta-rewrite-rules*) |
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
423 |
fun agoal atomic thy = agoalw atomic thy []; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
424 |
val goal = agoal false; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
425 |
|
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
426 |
(*now the versions that wrap the goal up in `Goal' to make it atomic*) |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
427 |
val atomic_goalw = agoalw true; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
428 |
val atomic_goal = agoal true; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4994
diff
changeset
|
429 |
|
6189 | 430 |
(*implicit context versions*) |
431 |
fun Goal s = atomic_goal (Context.the_context ()) s; |
|
432 |
fun Goalw thms s = atomic_goalw (Context.the_context ()) thms s; |
|
433 |
||
0 | 434 |
|
435 |
(*Proof step "by" the given tactic -- apply tactic to the proof state*) |
|
436 |
fun by_com tac ((th,ths), pairs) : gstack = |
|
4270 | 437 |
(case Seq.pull(tac th) of |
0 | 438 |
None => error"by: tactic failed" |
439 |
| Some(th2,ths2) => |
|
440 |
(if eq_thm(th,th2) |
|
3707
40856b593501
Prints warnings using the "warning" function instead of "writeln"
paulson
parents:
3669
diff
changeset
|
441 |
then warning "Warning: same as previous level" |
1460 | 442 |
else if eq_thm_sg(th,th2) then () |
4280 | 443 |
else warning ("Warning: signature of proof state has changed" ^ |
1460 | 444 |
sign_error (#sign(rep_thm th), #sign(rep_thm th2))); |
0 | 445 |
((th2,ths2)::(th,ths)::pairs))); |
446 |
||
8999 | 447 |
fun by tac = cond_timeit (!Library.timing) |
0 | 448 |
(fn() => make_command (by_com tac)); |
449 |
||
450 |
(* byev[tac1,...,tacn] applies tac1 THEN ... THEN tacn. |
|
451 |
Good for debugging proofs involving prove_goal.*) |
|
452 |
val byev = by o EVERY; |
|
453 |
||
454 |
||
455 |
(*Backtracking means find an alternative result from a tactic. |
|
456 |
If none at this level, try earlier levels*) |
|
457 |
fun backtrack [] = error"back: no alternatives" |
|
458 |
| backtrack ((th,thstr) :: pairs) = |
|
4270 | 459 |
(case Seq.pull thstr of |
1460 | 460 |
None => (writeln"Going back a level..."; backtrack pairs) |
461 |
| Some(th2,thstr2) => |
|
462 |
(if eq_thm(th,th2) |
|
3707
40856b593501
Prints warnings using the "warning" function instead of "writeln"
paulson
parents:
3669
diff
changeset
|
463 |
then warning "Warning: same as previous choice at this level" |
1460 | 464 |
else if eq_thm_sg(th,th2) then () |
3707
40856b593501
Prints warnings using the "warning" function instead of "writeln"
paulson
parents:
3669
diff
changeset
|
465 |
else warning "Warning: signature of proof state has changed"; |
1460 | 466 |
(th2,thstr2)::pairs)); |
0 | 467 |
|
468 |
fun back() = setstate (backtrack (getstate())); |
|
469 |
||
470 |
(*Chop back to previous level of the proof*) |
|
471 |
fun choplev n = make_command (chop_level n); |
|
472 |
||
473 |
(*Chopping back the goal stack*) |
|
474 |
fun chop () = make_command (fn (_,pairs) => pairs); |
|
475 |
||
476 |
(*Restore the previous proof state; discard current state. *) |
|
477 |
fun undo() = case !undo_list of |
|
478 |
[] => error"No proof state" |
|
479 |
| [_] => error"Already at initial state" |
|
480 |
| _::newundo => (undo_list := newundo; pr()) ; |
|
481 |
||
482 |
||
483 |
(*** Managing the proof stack ***) |
|
484 |
||
485 |
fun save_proof() = Proof(!undo_list, !curr_prems, !curr_mkresult); |
|
486 |
||
487 |
fun restore_proof(Proof(ul,prems,mk)) = |
|
488 |
(undo_list:= ul; curr_prems:= prems; curr_mkresult := mk; prems); |
|
489 |
||
490 |
||
491 |
fun top_proof() = case !proofstack of |
|
1460 | 492 |
[] => error("Stack of proof attempts is empty!") |
0 | 493 |
| p::ps => (p,ps); |
494 |
||
495 |
(* push a copy of the current proof state on to the stack *) |
|
496 |
fun push_proof() = (proofstack := (save_proof() :: !proofstack)); |
|
497 |
||
498 |
(* discard the top proof state of the stack *) |
|
499 |
fun pop_proof() = |
|
500 |
let val (p,ps) = top_proof() |
|
501 |
val prems = restore_proof p |
|
502 |
in proofstack := ps; pr(); prems end; |
|
503 |
||
504 |
(* rotate the stack so that the top element goes to the bottom *) |
|
505 |
fun rotate_proof() = let val (p,ps) = top_proof() |
|
1460 | 506 |
in proofstack := ps@[save_proof()]; |
507 |
restore_proof p; |
|
508 |
pr(); |
|
509 |
!curr_prems |
|
510 |
end; |
|
0 | 511 |
|
512 |
||
513 |
(** Shortcuts for commonly-used tactics **) |
|
514 |
||
515 |
fun bws rls = by (rewrite_goals_tac rls); |
|
516 |
fun bw rl = bws [rl]; |
|
517 |
||
518 |
fun brs rls i = by (resolve_tac rls i); |
|
519 |
fun br rl = brs [rl]; |
|
520 |
||
521 |
fun bes rls i = by (eresolve_tac rls i); |
|
522 |
fun be rl = bes [rl]; |
|
523 |
||
524 |
fun bds rls i = by (dresolve_tac rls i); |
|
525 |
fun bd rl = bds [rl]; |
|
526 |
||
527 |
fun ba i = by (assume_tac i); |
|
528 |
||
529 |
fun ren str i = by (rename_tac str i); |
|
530 |
||
531 |
(** Shortcuts to work on the first applicable subgoal **) |
|
532 |
||
533 |
fun frs rls = by (FIRSTGOAL (trace_goalno_tac (resolve_tac rls))); |
|
534 |
fun fr rl = frs [rl]; |
|
535 |
||
536 |
fun fes rls = by (FIRSTGOAL (trace_goalno_tac (eresolve_tac rls))); |
|
537 |
fun fe rl = fes [rl]; |
|
538 |
||
539 |
fun fds rls = by (FIRSTGOAL (trace_goalno_tac (dresolve_tac rls))); |
|
540 |
fun fd rl = fds [rl]; |
|
541 |
||
542 |
fun fa() = by (FIRSTGOAL (trace_goalno_tac assume_tac)); |
|
543 |
||
544 |
(** Reading and printing terms wrt the current theory **) |
|
545 |
||
546 |
fun top_sg() = #sign(rep_thm(topthm())); |
|
547 |
||
8086 | 548 |
fun read s = term_of (read_cterm (top_sg()) (s, TypeInfer.logicT)); |
0 | 549 |
|
550 |
(*Print a term under the current signature of the proof state*) |
|
551 |
fun prin t = writeln (Sign.string_of_term (top_sg()) t); |
|
552 |
||
553 |
fun printyp T = writeln (Sign.string_of_typ (top_sg()) T); |
|
554 |
||
555 |
fun pprint_term t = Sign.pprint_term (top_sg()) t; |
|
556 |
||
557 |
fun pprint_typ T = Sign.pprint_typ (top_sg()) T; |
|
558 |
||
1628
60136fdd80c4
Added functions pr_latex and printgoal_latex which
berghofe
parents:
1580
diff
changeset
|
559 |
|
0 | 560 |
(*Prints exceptions nicely at top level; |
561 |
raises the exception in order to have a polymorphic type!*) |
|
914
cae574c09137
Now prove_goalw_cterm calls print_sign_exn_unit, so that the rest
lcp
parents:
696
diff
changeset
|
562 |
fun print_exn e = (print_sign_exn_unit (top_sg()) e; raise e); |
0 | 563 |
|
11884 | 564 |
|
565 |
||
566 |
(** theorem database operations **) |
|
567 |
||
568 |
(* storing *) |
|
569 |
||
570 |
fun bind_thm (name, thm) = ThmDatabase.ml_store_thm (name, standard thm); |
|
571 |
fun bind_thms (name, thms) = ThmDatabase.ml_store_thms (name, map standard thms); |
|
572 |
||
573 |
fun qed name = ThmDatabase.ml_store_thm (name, result ()); |
|
574 |
fun qed_goal name thy goal tacsf = ThmDatabase.ml_store_thm (name, prove_goal thy goal tacsf); |
|
575 |
fun qed_goalw name thy rews goal tacsf = |
|
576 |
ThmDatabase.ml_store_thm (name, prove_goalw thy rews goal tacsf); |
|
577 |
fun qed_spec_mp name = |
|
578 |
ThmDatabase.ml_store_thm (name, ObjectLogic.rulify_no_asm (result ())); |
|
579 |
fun qed_goal_spec_mp name thy s p = |
|
580 |
bind_thm (name, ObjectLogic.rulify_no_asm (prove_goal thy s p)); |
|
581 |
fun qed_goalw_spec_mp name thy defs s p = |
|
582 |
bind_thm (name, ObjectLogic.rulify_no_asm (prove_goalw thy defs s p)); |
|
583 |
||
584 |
fun no_qed () = (); |
|
585 |
||
586 |
||
587 |
(* retrieval *) |
|
588 |
||
589 |
fun theory_of_goal () = ThyInfo.theory_of_sign (Thm.sign_of_thm (topthm ())); |
|
590 |
||
591 |
fun thms_containing raw_consts = |
|
592 |
let |
|
593 |
val thy = theory_of_goal (); |
|
594 |
val consts = map (Sign.intern_const (Theory.sign_of thy)) raw_consts; |
|
595 |
in ThmDatabase.thms_containing thy consts end; |
|
596 |
||
597 |
||
598 |
(*top_const: main constant, ignoring Trueprop*) |
|
599 |
fun top_const (_ $ t) = (case head_of t of Const (c, _) => Some c | _ => None) |
|
600 |
| top_const _ = None; |
|
601 |
||
602 |
val intro_const = top_const o concl_of; |
|
603 |
||
604 |
fun elim_const thm = case prems_of thm of [] => None | p::_ => top_const p; |
|
605 |
||
606 |
(* In case faster access is necessary, the thm db should provide special |
|
607 |
functions |
|
608 |
||
609 |
index_intros, index_elims: string -> (string * thm) list |
|
610 |
||
611 |
where thm [| A1 ; ...; An |] ==> B is returned by |
|
612 |
- index_intros c if B is of the form c t1 ... tn |
|
613 |
- index_elims c if A1 is of the form c t1 ... tn |
|
614 |
*) |
|
615 |
||
616 |
(* could be provided by thm db *) |
|
617 |
fun index_intros c = |
|
618 |
let fun topc(_,thm) = intro_const thm = Some(c); |
|
619 |
val named_thms = ThmDatabase.thms_containing (theory_of_goal ()) [c] |
|
620 |
in filter topc named_thms end; |
|
621 |
||
622 |
(* could be provided by thm db *) |
|
623 |
fun index_elims c = |
|
624 |
let fun topc(_,thm) = elim_const thm = Some(c); |
|
625 |
val named_thms = ThmDatabase.thms_containing (theory_of_goal ()) [c] |
|
626 |
in filter topc named_thms end; |
|
627 |
||
628 |
(*assume that parameters have unique names; reverse them for substitution*) |
|
629 |
fun goal_params i = |
|
630 |
let val gi = getgoal i |
|
631 |
val rfrees = rev(map Free (Logic.strip_params gi)) |
|
632 |
in (gi,rfrees) end; |
|
633 |
||
634 |
fun concl_of_goal i = |
|
635 |
let val (gi,rfrees) = goal_params i |
|
636 |
val B = Logic.strip_assums_concl gi |
|
637 |
in subst_bounds(rfrees,B) end; |
|
638 |
||
639 |
fun prems_of_goal i = |
|
640 |
let val (gi,rfrees) = goal_params i |
|
641 |
val As = Logic.strip_assums_hyp gi |
|
642 |
in map (fn A => subst_bounds(rfrees,A)) As end; |
|
643 |
||
644 |
(*returns all those named_thms whose subterm extracted by extract can be |
|
645 |
instantiated to obj; the list is sorted according to the number of premises |
|
646 |
and the size of the required substitution.*) |
|
647 |
fun select_match(obj, signobj, named_thms, extract) = |
|
648 |
let fun matches(prop, tsig) = |
|
649 |
case extract prop of |
|
650 |
None => false |
|
651 |
| Some pat => Pattern.matches tsig (pat, obj); |
|
652 |
||
653 |
fun substsize(prop, tsig) = |
|
654 |
let val Some pat = extract prop |
|
655 |
val (_,subst) = Pattern.match tsig (pat,obj) |
|
656 |
in foldl op+ |
|
657 |
(0, map (fn (_,t) => size_of_term t) subst) |
|
658 |
end |
|
659 |
||
660 |
fun thm_ord ((p0,s0,_),(p1,s1,_)) = |
|
661 |
prod_ord (int_ord o pairself (fn 0 => 0 | x => 1)) int_ord ((p0,s0),(p1,s1)); |
|
662 |
||
663 |
fun select((p as (_,thm))::named_thms, sels) = |
|
664 |
let val {prop, sign, ...} = rep_thm thm |
|
665 |
in select(named_thms, |
|
666 |
if Sign.subsig(sign, signobj) andalso |
|
667 |
matches(prop,#tsig(Sign.rep_sg signobj)) |
|
668 |
then |
|
669 |
(nprems_of thm,substsize(prop,#tsig(Sign.rep_sg signobj)),p)::sels |
|
670 |
else sels) |
|
671 |
end |
|
672 |
| select([],sels) = sels |
|
673 |
||
674 |
in map (fn (_,_,t) => t) (sort thm_ord (select(named_thms, []))) end; |
|
675 |
||
676 |
fun find_matching(prop,sign,index,extract) = |
|
677 |
(case top_const prop of |
|
678 |
Some c => select_match(prop,sign,index c,extract) |
|
679 |
| _ => []); |
|
680 |
||
681 |
fun find_intros(prop,sign) = |
|
682 |
find_matching(prop,sign,index_intros,Some o Logic.strip_imp_concl); |
|
683 |
||
684 |
fun find_elims sign prop = |
|
685 |
let fun major prop = case Logic.strip_imp_prems prop of |
|
686 |
[] => None | p::_ => Some p |
|
687 |
in find_matching(prop,sign,index_elims,major) end; |
|
688 |
||
689 |
fun findI i = find_intros(concl_of_goal i,#sign(rep_thm(topthm()))); |
|
690 |
||
691 |
fun findEs i = |
|
692 |
let fun eq_nth((n1,th1),(n2,th2)) = n1=n2 andalso eq_thm(th1,th2); |
|
693 |
val sign = #sign(rep_thm(topthm())) |
|
694 |
val thmss = map (find_elims sign) (prems_of_goal i) |
|
695 |
in foldl (gen_union eq_nth) ([],thmss) end; |
|
696 |
||
697 |
fun findE i j = |
|
698 |
let val sign = #sign(rep_thm(topthm())) |
|
699 |
in find_elims sign (nth_elem(j-1, prems_of_goal i)) end; |
|
700 |
||
0 | 701 |
end; |
1500 | 702 |
|
703 |
open Goals; |