author | wenzelm |
Wed, 12 Mar 2025 11:39:00 +0100 | |
changeset 82265 | 4b875a4c83b0 |
parent 80910 | 406a85a25189 |
permissions | -rw-r--r-- |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
1 |
(* Title: HOL/ex/Sketch_and_Explore.thy |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
2 |
Author: Florian Haftmann, TU Muenchen |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
3 |
*) |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
4 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
5 |
chapter \<open>Experimental commands \<^text>\<open>sketch\<close> and \<^text>\<open>explore\<close>\<close> |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
6 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
7 |
theory Sketch_and_Explore |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
8 |
imports Main \<comment> \<open>TODO: generalize existing sledgehammer functions to Pure\<close> |
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
9 |
keywords "sketch" "explore" "sketch_subgoals" :: diag |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
10 |
begin |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
11 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
12 |
ML \<open> |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
13 |
fun split_clause t = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
14 |
let |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
70225
diff
changeset
|
15 |
val (fixes, horn) = funpow_yield (length (Term.strip_all_vars t)) Logic.dest_all_global t; |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
16 |
val assms = Logic.strip_imp_prems horn; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
17 |
val concl = Logic.strip_imp_concl horn; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
18 |
in (fixes, assms, concl) end; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
19 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
20 |
fun print_typ ctxt T = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
21 |
T |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
22 |
|> Syntax.string_of_typ ctxt |
79824 | 23 |
|> ATP_Util.maybe_quote ctxt; |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
24 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
25 |
fun print_term ctxt t = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
26 |
t |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
27 |
|> singleton (Syntax.uncheck_terms ctxt) |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
28 |
|> Sledgehammer_Isar_Annotate.annotate_types_in_term ctxt |
80866
8c67b14fdd48
clarified internal tool output: prefer Pretty.pure_string_of over manipulation of print_mode;
wenzelm
parents:
80799
diff
changeset
|
29 |
|> Syntax.unparse_term ctxt |
8c67b14fdd48
clarified internal tool output: prefer Pretty.pure_string_of over manipulation of print_mode;
wenzelm
parents:
80799
diff
changeset
|
30 |
|> Pretty.pure_string_of |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
31 |
|> Sledgehammer_Util.simplify_spaces |
79824 | 32 |
|> ATP_Util.maybe_quote ctxt; |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
33 |
|
79902
0d5c41ea208a
sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents:
79901
diff
changeset
|
34 |
fun eigen_context_for_statement (params, assms, concl) ctxt = |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
35 |
let |
79902
0d5c41ea208a
sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents:
79901
diff
changeset
|
36 |
val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params |
0d5c41ea208a
sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents:
79901
diff
changeset
|
37 |
val ctxt' = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd |
80137
0c51e0a6bc37
sketch & explore: recover from duplicate fixed variables in Isar proofs
Simon Wimmer <wimmers@in.tum.de>
parents:
80100
diff
changeset
|
38 |
handle ERROR _ => |
0c51e0a6bc37
sketch & explore: recover from duplicate fixed variables in Isar proofs
Simon Wimmer <wimmers@in.tum.de>
parents:
80100
diff
changeset
|
39 |
ctxt |> Variable.set_body true |> Proof_Context.add_fixes fixes |> snd |
79902
0d5c41ea208a
sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents:
79901
diff
changeset
|
40 |
in ((params, assms, concl), ctxt') end; |
70225
129757af1096
more correct simulation of eigen context for generated Isar statements
haftmann
parents:
70018
diff
changeset
|
41 |
|
129757af1096
more correct simulation of eigen context for generated Isar statements
haftmann
parents:
70018
diff
changeset
|
42 |
fun print_isar_skeleton ctxt indent keyword stmt = |
129757af1096
more correct simulation of eigen context for generated Isar statements
haftmann
parents:
70018
diff
changeset
|
43 |
let |
129757af1096
more correct simulation of eigen context for generated Isar statements
haftmann
parents:
70018
diff
changeset
|
44 |
val ((fixes, assms, concl), ctxt') = eigen_context_for_statement stmt ctxt; |
80799 | 45 |
val prefix = Symbol.spaces indent; |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
46 |
val prefix_sep = "\n" ^ prefix ^ " and "; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
47 |
val show_s = prefix ^ keyword ^ " " ^ print_term ctxt' concl; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
48 |
val if_s = if null assms then NONE |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
49 |
else SOME (prefix ^ " if " ^ space_implode prefix_sep |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
50 |
(map (print_term ctxt') assms)); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
51 |
val for_s = if null fixes then NONE |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
52 |
else SOME (prefix ^ " for " ^ space_implode prefix_sep |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
53 |
(map (fn (v, T) => v ^ " :: " ^ print_typ ctxt T) fixes)); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
54 |
val s = cat_lines ([show_s] @ map_filter I [if_s, for_s] @ |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
55 |
[prefix ^ " " ^ (if is_none if_s then "" else "using that ") ^ "sorry"]); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
56 |
in |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
57 |
s |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
58 |
end; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
59 |
|
79901
54e9875e491f
sketch & explore: replace functionality of `sketch` by more useful `nxsketch` (and remove `nxsketch`)
Simon Wimmer <wimmers@in.tum.de>
parents:
79900
diff
changeset
|
60 |
fun print_skeleton ctxt indent keyword stmt = |
78820
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
61 |
let |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
62 |
val ((fixes, assms, concl), ctxt') = eigen_context_for_statement stmt ctxt; |
80799 | 63 |
val prefix = Symbol.spaces indent; |
79900
8f0ff2847ba8
sketch & explore: use Active.sendback_markup_command to preserve indentation of generated proof text
Simon Wimmer <wimmers@in.tum.de>
parents:
79824
diff
changeset
|
64 |
val prefix_sep = "\n" ^ prefix ^ " and "; |
78820
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
65 |
val show_s = prefix ^ keyword ^ " " ^ print_term ctxt' concl; |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
66 |
val assumes_s = if null assms then NONE |
79900
8f0ff2847ba8
sketch & explore: use Active.sendback_markup_command to preserve indentation of generated proof text
Simon Wimmer <wimmers@in.tum.de>
parents:
79824
diff
changeset
|
67 |
else SOME (prefix ^ "assume " ^ space_implode prefix_sep |
78820
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
68 |
(map (print_term ctxt') assms)); |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
69 |
val fixes_s = if null fixes then NONE |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
70 |
else SOME (prefix ^ "fix " ^ space_implode prefix_sep |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
71 |
(map (fn (v, T) => v ^ " :: " ^ print_typ ctxt T) fixes)); |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
72 |
val s = cat_lines (map_filter I [fixes_s, assumes_s] @ [show_s] @ [prefix ^ " sorry"]); |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
73 |
in |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
74 |
s |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
75 |
end; |
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
76 |
|
79901
54e9875e491f
sketch & explore: replace functionality of `sketch` by more useful `nxsketch` (and remove `nxsketch`)
Simon Wimmer <wimmers@in.tum.de>
parents:
79900
diff
changeset
|
77 |
fun print_sketch ctxt method_text clauses = |
54e9875e491f
sketch & explore: replace functionality of `sketch` by more useful `nxsketch` (and remove `nxsketch`)
Simon Wimmer <wimmers@in.tum.de>
parents:
79900
diff
changeset
|
78 |
"proof" ^ method_text :: separate "next" (map (print_skeleton ctxt 2 "show") clauses) @ ["qed"]; |
78820
b356019e8d49
Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents:
74525
diff
changeset
|
79 |
|
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
80 |
fun print_exploration ctxt method_text [clause] = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
81 |
["proof -", print_isar_skeleton ctxt 2 "have" clause, |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
82 |
" then show ?thesis", " by" ^ method_text, "qed"] |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
83 |
| print_exploration ctxt method_text (clause :: clauses) = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
84 |
"proof -" :: print_isar_skeleton ctxt 2 "have" clause |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
85 |
:: map (print_isar_skeleton ctxt 2 "moreover have") clauses |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
86 |
@ [" ultimately show ?thesis", " by" ^ method_text, "qed"]; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
87 |
|
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
88 |
fun print_subgoal apply_line_opt (is_prems, is_for, is_sh) ctxt indent stmt = |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
89 |
let |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
90 |
val ((fixes, _, _), ctxt') = eigen_context_for_statement stmt ctxt; |
80799 | 91 |
val prefix = Symbol.spaces indent; |
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
92 |
val fixes_s = |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
93 |
if not is_for orelse null fixes then NONE |
80910 | 94 |
else SOME ("for " ^ implode_space |
80307 | 95 |
(map (fn (v, _) => Thy_Header.print_name' ctxt' v) fixes)); |
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
96 |
val premises_s = if is_prems then SOME "premises prems" else NONE; |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
97 |
val sh_s = if is_sh then SOME "sledgehammer" else NONE; |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
98 |
val subgoal_s = map_filter I [SOME "subgoal", premises_s, fixes_s] |
80910 | 99 |
|> implode_space; |
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
100 |
val using_s = if is_prems then SOME "using prems" else NONE; |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
101 |
val s = cat_lines ( |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
102 |
[subgoal_s] |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
103 |
@ map_filter (Option.map (fn s => prefix ^ s)) [using_s, apply_line_opt, sh_s] |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
104 |
@ [prefix ^ "sorry"]); |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
105 |
in |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
106 |
s |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
107 |
end; |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
108 |
|
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
109 |
fun coalesce_method_txt [] = "" |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
110 |
| coalesce_method_txt [s] = s |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
111 |
| coalesce_method_txt (s1 :: s2 :: ss) = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
112 |
if s1 = "(" orelse s1 = "[" |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
113 |
orelse s2 = ")" orelse s2 = "]" orelse s2= ":" |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
114 |
then s1 ^ coalesce_method_txt (s2 :: ss) |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
115 |
else s1 ^ " " ^ coalesce_method_txt (s2 :: ss); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
116 |
|
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
117 |
fun print_subgoals options apply_line_opt ctxt _ clauses = |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
118 |
separate "" (map (print_subgoal apply_line_opt options ctxt 2) clauses) @ ["done"]; |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
119 |
|
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
120 |
fun print_proof_text_from_state print (some_method_ref : ((Method.text * Position.range) * Token.T list) option) state = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
121 |
let |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
122 |
val state' = state |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
123 |
|> Proof.proof (Option.map fst some_method_ref) |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
124 |
|> Seq.the_result "" |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
125 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
126 |
val { context = ctxt, facts = _, goal } = Proof.goal state'; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
127 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
128 |
val ctxt_print = fold (fn opt => Config.put opt false) |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
129 |
[show_markup, Printer.show_type_emphasis, show_types, show_sorts, show_consts] ctxt |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
130 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
131 |
val method_text = case some_method_ref of |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
132 |
NONE => "" |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
133 |
| SOME (_, toks) => " " ^ coalesce_method_txt (map Token.unparse toks); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
134 |
\<comment> \<open>TODO proper printing required\<close> |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
135 |
val goal_props = Logic.strip_imp_prems (Thm.prop_of goal); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
136 |
val clauses = map split_clause goal_props; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
137 |
val lines = if null clauses then |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
138 |
if is_none some_method_ref then [" .."] |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
139 |
else [" by" ^ method_text] |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
140 |
else print ctxt_print method_text clauses; |
80308 | 141 |
in Output.information (Active.sendback_markup_command (cat_lines lines)) end; |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
142 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
143 |
val sketch = print_proof_text_from_state print_sketch; |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
144 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
145 |
fun explore method_ref = print_proof_text_from_state print_exploration (SOME method_ref); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
146 |
|
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
147 |
fun subgoals options method_ref = |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
148 |
let |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
149 |
val apply_line_opt = case method_ref of |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
150 |
NONE => NONE |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
151 |
| SOME (_, toks) => SOME ("apply " ^ coalesce_method_txt (map Token.unparse toks)) |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
152 |
val succeed_method_ref = SOME ((Method.succeed_text, Position.no_range), []) |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
153 |
in |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
154 |
print_proof_text_from_state (print_subgoals options apply_line_opt) succeed_method_ref |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
155 |
end; |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
156 |
|
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
157 |
fun sketch_cmd some_method_text = |
80308 | 158 |
Toplevel.keep_proof (fn state => sketch some_method_text (Toplevel.proof_of state)); |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
159 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
160 |
fun explore_cmd method_text = |
80308 | 161 |
Toplevel.keep_proof (fn state => explore method_text (Toplevel.proof_of state)); |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
162 |
|
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
163 |
fun subgoals_cmd (modes, method_ref) = |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
164 |
let |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
165 |
val is_prems = not (member (op =) modes "noprems") |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
166 |
val is_for = not (member (op =) modes "nofor") |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
167 |
val is_sh = member (op =) modes "sh" |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
168 |
in |
80308 | 169 |
Toplevel.keep_proof (fn state => |
170 |
subgoals (is_prems, is_for, is_sh) method_ref (Toplevel.proof_of state)) |
|
171 |
end; |
|
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
172 |
|
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
173 |
val _ = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
174 |
Outer_Syntax.command \<^command_keyword>\<open>sketch\<close> |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
175 |
"print sketch of Isar proof text after method application" |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
176 |
(Scan.option (Scan.trace Method.parse) >> sketch_cmd); |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
177 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
178 |
val _ = |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
179 |
Outer_Syntax.command \<^command_keyword>\<open>explore\<close> |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
180 |
"explore proof obligations after method application as Isar proof text" |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
181 |
(Scan.trace Method.parse >> explore_cmd); |
80100
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
182 |
|
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
183 |
val opt_modes = |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
184 |
Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) []; |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
185 |
|
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
186 |
val _ = |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
187 |
Outer_Syntax.command \<^command_keyword>\<open>sketch_subgoals\<close> |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
188 |
"sketch proof obligations as subgoals, applying a method and/or sledgehammer to each" |
7506cb70ecfb
Add subgoals variant of 'sketch' command
Simon Wimmer <wimmers@in.tum.de>
parents:
79903
diff
changeset
|
189 |
(opt_modes -- Scan.option (Scan.trace Method.parse) >> subgoals_cmd); |
70018
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
190 |
\<close> |
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
191 |
|
571909ef3103
experimental commands for proof sketching and exploration
haftmann
parents:
diff
changeset
|
192 |
end |