author | wenzelm |
Sun, 05 Jul 2015 15:02:30 +0200 | |
changeset 60642 | 48dd1cefb4ae |
parent 60555 | 51a6997b1384 |
child 60695 | 757549b4bbe6 |
permissions | -rw-r--r-- |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/Isar/obtain.ML |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
2 |
Author: Markus Wenzel, TU Muenchen |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
3 |
|
60452 | 4 |
Generalized existence and cases rules within Isar proof text. |
8094 | 5 |
*) |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
6 |
|
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
7 |
signature OBTAIN = |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
8 |
sig |
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
9 |
val obtain_thesis: Proof.context -> ((string * typ) * term) * Proof.context |
60444 | 10 |
val obtains_attributes: ('typ, 'term) Element.obtain list -> attribute list |
60448 | 11 |
val read_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list |
12 |
val cert_obtains: Proof.context -> term -> Element.obtains_i -> (binding * term) list |
|
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
13 |
val parse_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list |
60448 | 14 |
val consider: Element.obtains_i -> bool -> Proof.state -> Proof.state |
15 |
val consider_cmd: Element.obtains -> bool -> Proof.state -> Proof.state |
|
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
16 |
val obtain: binding -> (binding * typ option * mixfix) list -> |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
17 |
(Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state |
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
18 |
val obtain_cmd: binding -> (binding * string option * mixfix) list -> |
30211 | 19 |
(Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state |
20308 | 20 |
val result: (Proof.context -> tactic) -> thm list -> Proof.context -> |
32199 | 21 |
((string * cterm) list * thm list) * Proof.context |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
22 |
val guess: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state |
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
23 |
val guess_cmd: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
24 |
end; |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
25 |
|
10379
93630e0c5ae9
improved handling of "that": insert into goal, only declare as Pure "intro";
wenzelm
parents:
9481
diff
changeset
|
26 |
structure Obtain: OBTAIN = |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
27 |
struct |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
28 |
|
60452 | 29 |
(** specification elements **) |
30 |
||
31 |
(* obtain_export *) |
|
18670 | 32 |
|
18870 | 33 |
(* |
18897 | 34 |
[x, A x] |
35 |
: |
|
36 |
B |
|
37 |
-------- |
|
38 |
B |
|
18870 | 39 |
*) |
21686 | 40 |
fun eliminate_term ctxt xs tm = |
41 |
let |
|
42 |
val vs = map (dest_Free o Thm.term_of) xs; |
|
43 |
val bads = Term.fold_aterms (fn t as Free v => |
|
44 |
if member (op =) vs v then insert (op aconv) t else I | _ => I) tm []; |
|
45 |
val _ = null bads orelse |
|
46 |
error ("Result contains obtained parameters: " ^ |
|
24920 | 47 |
space_implode " " (map (Syntax.string_of_term ctxt) bads)); |
21686 | 48 |
in tm end; |
49 |
||
60387
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
50 |
fun eliminate ctxt rule xs As thm = |
9468 | 51 |
let |
60387
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
52 |
val _ = eliminate_term ctxt xs (Thm.full_prop_of thm); |
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
53 |
val _ = Object_Logic.is_judgment ctxt (Thm.concl_of thm) orelse |
20308 | 54 |
error "Conclusion in obtained context must be object-logic judgment"; |
55 |
||
60387
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
56 |
val ((_, [thm']), ctxt') = Variable.import true [thm] ctxt; |
20308 | 57 |
val prems = Drule.strip_imp_prems (#prop (Thm.crep_thm thm')); |
9468 | 58 |
in |
20308 | 59 |
((Drule.implies_elim_list thm' (map Thm.assume prems) |
60315 | 60 |
|> Drule.implies_intr_list (map (Drule.norm_hhf_cterm ctxt') As) |
20308 | 61 |
|> Drule.forall_intr_list xs) |
62 |
COMP rule) |
|
63 |
|> Drule.implies_intr_list prems |
|
60387
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
64 |
|> singleton (Variable.export ctxt' ctxt) |
9468 | 65 |
end; |
66 |
||
21686 | 67 |
fun obtain_export ctxt rule xs _ As = |
68 |
(eliminate ctxt rule xs As, eliminate_term ctxt xs); |
|
69 |
||
9468 | 70 |
|
60448 | 71 |
(* result declaration *) |
72 |
||
60444 | 73 |
fun obtains_attributes (obtains: ('typ, 'term) Element.obtain list) = |
74 |
let |
|
75 |
val case_names = obtains |> map_index (fn (i, (b, _)) => |
|
76 |
if Binding.is_empty b then string_of_int (i + 1) else Name_Space.base_name b); |
|
77 |
in [Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names case_names] end; |
|
78 |
||
79 |
||
60448 | 80 |
(* obtain thesis *) |
81 |
||
82 |
fun obtain_thesis ctxt = |
|
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
83 |
let |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
84 |
val ([x], ctxt') = |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
85 |
Proof_Context.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] ctxt; |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
86 |
val t = Object_Logic.fixed_judgment ctxt x; |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
87 |
val v = dest_Free (Object_Logic.drop_judgment ctxt t); |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
88 |
in ((v, t), ctxt') end; |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
89 |
|
60448 | 90 |
|
91 |
(* obtain clauses *) |
|
92 |
||
93 |
local |
|
94 |
||
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
95 |
val mk_all_external = Logic.all_constraint o Variable.default_type; |
60481 | 96 |
|
97 |
fun mk_all_internal ctxt (y, z) t = |
|
98 |
let |
|
99 |
val T = |
|
100 |
(case AList.lookup (op =) (Term.add_frees t []) z of |
|
101 |
SOME T => T |
|
102 |
| NONE => the_default dummyT (Variable.default_type ctxt z)); |
|
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
103 |
in Logic.all_const T $ Term.lambda_name (y, Free (z, T)) t end; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
104 |
|
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
105 |
fun prepare_clause prep_var parse_prop mk_all ctxt thesis raw_vars raw_props = |
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
106 |
let |
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
107 |
val ((xs', vars), ctxt') = ctxt |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
108 |
|> fold_map prep_var raw_vars |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
109 |
|-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars); |
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
110 |
val xs = map (Variable.check_name o #1) vars; |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
111 |
in |
60448 | 112 |
Logic.list_implies (map (parse_prop ctxt') raw_props, thesis) |
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
113 |
|> fold_rev (mk_all ctxt') (xs ~~ xs') |
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
114 |
end; |
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
115 |
|
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
116 |
fun prepare_obtains prep_clause check_terms |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
117 |
ctxt thesis (raw_obtains: ('typ, 'term) Element.obtain list) = |
60448 | 118 |
let |
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
119 |
val clauses = raw_obtains |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
120 |
|> map (fn (_, (raw_vars, raw_props)) => prep_clause ctxt thesis raw_vars raw_props) |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
121 |
|> check_terms ctxt; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
122 |
in map fst raw_obtains ~~ clauses end; |
60448 | 123 |
|
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
124 |
val parse_clause = prepare_clause Proof_Context.read_var Syntax.parse_prop mk_all_external; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
125 |
val cert_clause = prepare_clause Proof_Context.cert_var (K I) mk_all_internal; |
60448 | 126 |
|
127 |
in |
|
128 |
||
60477
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
129 |
val read_obtains = prepare_obtains parse_clause Syntax.check_terms; |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
130 |
val cert_obtains = prepare_obtains cert_clause (K I); |
051b200f7578
improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents:
60475
diff
changeset
|
131 |
val parse_obtains = prepare_obtains parse_clause (K I); |
60448 | 132 |
|
133 |
end; |
|
134 |
||
135 |
||
136 |
||
60451
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
137 |
(** consider: generalized elimination and cases rule **) |
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
138 |
|
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
139 |
(* |
60475 | 140 |
consider (a) x where "A x" | (b) y where "B y" | ... == |
60451
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
141 |
|
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
142 |
have thesis |
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
143 |
if a [intro?]: "!!x. A x ==> thesis" |
60475 | 144 |
and b [intro?]: "!!y. B y ==> thesis" |
60451
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
145 |
and ... |
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
146 |
for thesis |
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
147 |
apply (insert that) |
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
148 |
*) |
60448 | 149 |
|
150 |
local |
|
151 |
||
152 |
fun gen_consider prep_obtains raw_obtains int state = |
|
153 |
let |
|
154 |
val _ = Proof.assert_forward_or_chain state; |
|
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
155 |
val ctxt = Proof.context_of state; |
60451
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
156 |
|
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
157 |
val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt; |
60448 | 158 |
val obtains = prep_obtains thesis_ctxt thesis raw_obtains; |
60456 | 159 |
val atts = Rule_Cases.cases_open :: obtains_attributes raw_obtains; |
60448 | 160 |
in |
161 |
state |
|
60555
51a6997b1384
support 'when' statement, which corresponds to 'presume';
wenzelm
parents:
60481
diff
changeset
|
162 |
|> Proof.have true NONE (K I) |
60451
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
163 |
[(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] |
1f2b29f78439
clarified 'consider', using structured 'have' statement;
wenzelm
parents:
60448
diff
changeset
|
164 |
(map (fn (a, A) => ((a, [Context_Rules.intro_query NONE]), [(A, [])])) obtains) |
60456 | 165 |
[((Binding.empty, atts), [(thesis, [])])] int |
60461 | 166 |
|-> Proof.refine_insert |
60448 | 167 |
end; |
168 |
||
169 |
in |
|
170 |
||
171 |
val consider = gen_consider cert_obtains; |
|
172 |
val consider_cmd = gen_consider read_obtains; |
|
173 |
||
174 |
end; |
|
175 |
||
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
176 |
|
60444 | 177 |
|
60452 | 178 |
(** obtain: augmented context based on generalized existence rule **) |
179 |
||
180 |
(* |
|
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
181 |
obtain (a) x where "A x" <proof> == |
60452 | 182 |
|
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
183 |
have thesis if a [intro?]: "!!x. A x ==> thesis" for thesis |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
184 |
apply (insert that) |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
185 |
<proof> |
60452 | 186 |
fix x assm <<obtain_export>> "A x" |
187 |
*) |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
188 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
189 |
local |
8094 | 190 |
|
60379 | 191 |
fun gen_obtain prep_att prep_var prep_propp |
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
192 |
that_binding raw_vars raw_asms int state = |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
193 |
let |
9468 | 194 |
val _ = Proof.assert_forward_or_chain state; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
195 |
val ctxt = Proof.context_of state; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
196 |
|
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
197 |
(*vars*) |
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
198 |
val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt; |
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
199 |
val ((xs', vars), fix_ctxt) = thesis_ctxt |
60392 | 200 |
|> fold_map prep_var raw_vars |
201 |
|-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars); |
|
42494 | 202 |
val xs = map (Variable.check_name o #1) vars; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
203 |
|
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
204 |
(*asms*) |
60408
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
205 |
val (propss, binds) = prep_propp fix_ctxt (map snd raw_asms); |
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
206 |
val props = flat propss; |
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
207 |
val declare_asms = |
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
208 |
fold Variable.declare_term props #> |
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
209 |
fold Variable.bind_term binds; |
60387
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
210 |
val asms = |
60392 | 211 |
map (fn ((b, raw_atts), _) => (b, map (prep_att fix_ctxt) raw_atts)) raw_asms ~~ |
60408
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
212 |
map (map (rpair [])) propss; |
10464 | 213 |
|
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
214 |
(*params*) |
60407 | 215 |
val (params, params_ctxt) = |
216 |
declare_asms fix_ctxt |> fold_map Proof_Context.inferred_param xs' |>> map Free; |
|
60387
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
217 |
val cparams = map (Thm.cterm_of params_ctxt) params; |
60408
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
218 |
val binds' = (map o apsnd) (fold_rev Term.dependent_lambda_name (xs ~~ params)) binds; |
60387
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents:
60383
diff
changeset
|
219 |
|
60401 | 220 |
val _ = Variable.warn_extra_tfrees fix_ctxt params_ctxt; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
221 |
|
60446
64f48e7f921f
support to parse obtain clause without type-checking yet;
wenzelm
parents:
60444
diff
changeset
|
222 |
(*statements*) |
10582 | 223 |
val that_prop = |
45328 | 224 |
Logic.list_rename_params xs |
60408
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents:
60407
diff
changeset
|
225 |
(fold_rev Logic.all params (Logic.list_implies (props, thesis))); |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
226 |
|
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
227 |
fun after_qed (result_ctxt, results) state' = |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
228 |
let val [rule] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
229 |
state' |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
230 |
|> Proof.fix vars |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
231 |
|> Proof.map_context declare_asms |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
232 |
|> Proof.assm (obtain_export params_ctxt rule cparams) asms |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
233 |
end; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
234 |
in |
8094 | 235 |
state |
60555
51a6997b1384
support 'when' statement, which corresponds to 'presume';
wenzelm
parents:
60481
diff
changeset
|
236 |
|> Proof.have true NONE after_qed |
60453
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
237 |
[(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
238 |
[((that_binding, [Context_Rules.intro_query NONE]), [(that_prop, [])])] |
ba9085f7433f
clarified 'obtain', using structured 'have' statement;
wenzelm
parents:
60452
diff
changeset
|
239 |
[(Thm.empty_binding, [(thesis, [])])] int |
60461 | 240 |
|-> Proof.refine_insert |
60401 | 241 |
|> Proof.map_context (fold Variable.bind_term binds') |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
242 |
end; |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
243 |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
244 |
in |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
245 |
|
60388 | 246 |
val obtain = gen_obtain (K I) Proof_Context.cert_var Proof_Context.cert_propp; |
247 |
val obtain_cmd = gen_obtain Attrib.attribute_cmd Proof_Context.read_var Proof_Context.read_propp; |
|
8094 | 248 |
|
249 |
end; |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
250 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
251 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
252 |
|
20308 | 253 |
(** tactical result **) |
254 |
||
255 |
fun check_result ctxt thesis th = |
|
256 |
(case Thm.prems_of th of |
|
257 |
[prem] => |
|
258 |
if Thm.concl_of th aconv thesis andalso |
|
259 |
Logic.strip_assums_concl prem aconv thesis then th |
|
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
31794
diff
changeset
|
260 |
else error ("Guessed a different clause:\n" ^ Display.string_of_thm ctxt th) |
38875
c7a66b584147
tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents:
36323
diff
changeset
|
261 |
| [] => error "Goal solved -- nothing guessed" |
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
31794
diff
changeset
|
262 |
| _ => error ("Guess split into several cases:\n" ^ Display.string_of_thm ctxt th)); |
20308 | 263 |
|
264 |
fun result tac facts ctxt = |
|
265 |
let |
|
60448 | 266 |
val ((thesis_var, thesis), thesis_ctxt) = obtain_thesis ctxt; |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59616
diff
changeset
|
267 |
val st = Goal.init (Thm.cterm_of ctxt thesis); |
20308 | 268 |
val rule = |
59616 | 269 |
(case SINGLE (Method.insert_tac facts 1 THEN tac thesis_ctxt) st of |
20308 | 270 |
NONE => raise THM ("Obtain.result: tactic failed", 0, facts) |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
52456
diff
changeset
|
271 |
| SOME th => |
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
52456
diff
changeset
|
272 |
check_result thesis_ctxt thesis (Raw_Simplifier.norm_hhf thesis_ctxt (Goal.conclude th))); |
20308 | 273 |
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59616
diff
changeset
|
274 |
val closed_rule = Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) rule; |
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30763
diff
changeset
|
275 |
val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt; |
59616 | 276 |
val obtain_rule = |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59616
diff
changeset
|
277 |
Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule'; |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
278 |
val ((params, stmt), fix_ctxt) = Variable.focus_cterm (Thm.cprem_of obtain_rule 1) ctxt'; |
20308 | 279 |
val (prems, ctxt'') = |
32199 | 280 |
Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params)) |
20308 | 281 |
(Drule.strip_imp_prems stmt) fix_ctxt; |
282 |
in ((params, prems), ctxt'') end; |
|
283 |
||
284 |
||
285 |
||
60452 | 286 |
(** guess: obtain based on tactical result **) |
287 |
||
288 |
(* |
|
289 |
<chain_facts> |
|
290 |
guess x <proof body> <proof end> == |
|
291 |
||
292 |
{ |
|
293 |
fix thesis |
|
294 |
<chain_facts> have "PROP ?guess" |
|
60475 | 295 |
apply magic -- {* turn goal into "thesis ==> #thesis" *} |
60452 | 296 |
<proof body> |
60475 | 297 |
apply_end magic -- {* turn final "(!!x. P x ==> thesis) ==> #thesis" into |
60452 | 298 |
"#((!!x. A x ==> thesis) ==> thesis)" which is a finished goal state *} |
299 |
<proof end> |
|
300 |
} |
|
301 |
fix x assm <<obtain_export>> "A x" |
|
302 |
*) |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
303 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
304 |
local |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
305 |
|
20308 | 306 |
fun unify_params vars thesis_var raw_rule ctxt = |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
307 |
let |
42360 | 308 |
val thy = Proof_Context.theory_of ctxt; |
39134
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
wenzelm
parents:
38875
diff
changeset
|
309 |
val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); |
17891 | 310 |
|
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
31794
diff
changeset
|
311 |
fun err msg th = error (msg ^ ":\n" ^ Display.string_of_thm ctxt th); |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
312 |
|
19978 | 313 |
val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1; |
19779 | 314 |
val rule = Thm.incr_indexes (maxidx + 1) raw_rule; |
315 |
||
33368 | 316 |
val params = Rule_Cases.strip_params (Logic.nth_prem (1, Thm.prop_of rule)); |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
317 |
val m = length vars; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
318 |
val n = length params; |
19779 | 319 |
val _ = m <= n orelse err "More variables than parameters in obtained rule" rule; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
320 |
|
19779 | 321 |
fun unify ((x, T), (y, U)) (tyenv, max) = Sign.typ_unify thy (T, U) (tyenv, max) |
322 |
handle Type.TUNIFY => |
|
323 |
err ("Failed to unify variable " ^ |
|
324 |
string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^ |
|
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
47815
diff
changeset
|
325 |
string_of_term (Syntax_Trans.mark_bound_abs (y, Envir.norm_type tyenv U)) ^ " in") rule; |
33957 | 326 |
val (tyenv, _) = fold unify (map #1 vars ~~ take m params) |
19779 | 327 |
(Vartab.empty, Int.max (maxidx, Thm.maxidx_of rule)); |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
328 |
val norm_type = Envir.norm_type tyenv; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
329 |
|
19978 | 330 |
val xs = map (apsnd norm_type o fst) vars; |
33957 | 331 |
val ys = map (apsnd norm_type) (drop m params); |
20085
c5d60752587f
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
20004
diff
changeset
|
332 |
val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys; |
59623 | 333 |
val terms = map (Drule.mk_term o Thm.cterm_of ctxt o Free) (xs @ ys'); |
19779 | 334 |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
335 |
val instT = |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
336 |
fold (Term.add_tvarsT o #2) params [] |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60555
diff
changeset
|
337 |
|> map (fn v => (v, Thm.ctyp_of ctxt (norm_type (TVar v)))); |
20308 | 338 |
val closed_rule = rule |
59623 | 339 |
|> Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) |
20308 | 340 |
|> Thm.instantiate (instT, []); |
17891 | 341 |
|
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30763
diff
changeset
|
342 |
val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt; |
19978 | 343 |
val vars' = |
344 |
map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~ |
|
345 |
(map snd vars @ replicate (length ys) NoSyn); |
|
59623 | 346 |
val rule'' = Thm.forall_elim (Thm.cterm_of ctxt' (Logic.varify_global (Free thesis_var))) rule'; |
19978 | 347 |
in ((vars', rule''), ctxt') end; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
348 |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
24920
diff
changeset
|
349 |
fun inferred_type (binding, _, mx) ctxt = |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
24920
diff
changeset
|
350 |
let |
42494 | 351 |
val x = Variable.check_name binding; |
60407 | 352 |
val ((_, T), ctxt') = Proof_Context.inferred_param x ctxt |
19779 | 353 |
in ((x, T, mx), ctxt') end; |
354 |
||
20004
e6d3f2b031e6
guess: proper context for polymorphic parameters;
wenzelm
parents:
19978
diff
changeset
|
355 |
fun polymorphic ctxt vars = |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19844
diff
changeset
|
356 |
let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars)) |
19779 | 357 |
in map2 (fn (x, _, mx) => fn T => ((x, T), mx)) vars Ts end; |
18693 | 358 |
|
60379 | 359 |
fun gen_guess prep_var raw_vars int state = |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
360 |
let |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
361 |
val _ = Proof.assert_forward_or_chain state; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
362 |
val ctxt = Proof.context_of state; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
363 |
val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else []; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
364 |
|
60448 | 365 |
val (thesis_var, thesis) = #1 (obtain_thesis ctxt); |
60379 | 366 |
val vars = ctxt |
367 |
|> fold_map prep_var raw_vars |-> fold_map inferred_type |
|
368 |
|> fst |> polymorphic ctxt; |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
369 |
|
19978 | 370 |
fun guess_context raw_rule state' = |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
371 |
let |
19978 | 372 |
val ((parms, rule), ctxt') = |
20308 | 373 |
unify_params vars thesis_var raw_rule (Proof.context_of state'); |
42501
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
374 |
val (xs, _) = Variable.add_fixes (map (#1 o #1) parms) ctxt'; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
375 |
val ps = xs ~~ map (#2 o #1) parms; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
376 |
val ts = map Free ps; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
377 |
val asms = |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
378 |
Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule)) |
46219
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
46215
diff
changeset
|
379 |
|> map (fn asm => (Term.betapplys (fold_rev Term.abs ps asm, ts), [])); |
19779 | 380 |
val _ = not (null asms) orelse error "Trivial result -- nothing guessed"; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
381 |
in |
19978 | 382 |
state' |
383 |
|> Proof.map_context (K ctxt') |
|
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
384 |
|> Proof.fix (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms) |
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
385 |
|> `Proof.context_of |-> (fn fix_ctxt => Proof.assm |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59616
diff
changeset
|
386 |
(obtain_export fix_ctxt rule (map (Thm.cterm_of ctxt) ts)) |
59616 | 387 |
[(Thm.empty_binding, asms)]) |
60401 | 388 |
|> Proof.map_context (fold Variable.unbind_term Auto_Bind.no_facts) |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
389 |
end; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
390 |
|
19779 | 391 |
val goal = Var (("guess", 0), propT); |
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
392 |
val pos = Position.thread_data (); |
19779 | 393 |
fun print_result ctxt' (k, [(s, [_, th])]) = |
56932
11a4001b06c6
more position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents:
56897
diff
changeset
|
394 |
Proof_Display.print_results int pos ctxt' (k, [(s, [th])]); |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
52456
diff
changeset
|
395 |
val before_qed = |
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
52456
diff
changeset
|
396 |
Method.primitive_text (fn ctxt => |
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
52456
diff
changeset
|
397 |
Goal.conclude #> Raw_Simplifier.norm_hhf ctxt #> |
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
52456
diff
changeset
|
398 |
(fn th => Goal.protect 0 (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th))); |
60415
9d37b2330ee3
clarified local after_qed: result is not exported yet;
wenzelm
parents:
60414
diff
changeset
|
399 |
fun after_qed (result_ctxt, results) state' = |
9d37b2330ee3
clarified local after_qed: result is not exported yet;
wenzelm
parents:
60414
diff
changeset
|
400 |
let val [_, res] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) |
9d37b2330ee3
clarified local after_qed: result is not exported yet;
wenzelm
parents:
60414
diff
changeset
|
401 |
in |
9d37b2330ee3
clarified local after_qed: result is not exported yet;
wenzelm
parents:
60414
diff
changeset
|
402 |
state' |
9d37b2330ee3
clarified local after_qed: result is not exported yet;
wenzelm
parents:
60414
diff
changeset
|
403 |
|> Proof.end_block |
9d37b2330ee3
clarified local after_qed: result is not exported yet;
wenzelm
parents:
60414
diff
changeset
|
404 |
|> guess_context (check_result ctxt thesis res) |
9d37b2330ee3
clarified local after_qed: result is not exported yet;
wenzelm
parents:
60414
diff
changeset
|
405 |
end; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
406 |
in |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
407 |
state |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
408 |
|> Proof.enter_forward |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
409 |
|> Proof.begin_block |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
410 |
|> Proof.fix [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
411 |
|> Proof.chain_facts chain_facts |
60555
51a6997b1384
support 'when' statement, which corresponds to 'presume';
wenzelm
parents:
60481
diff
changeset
|
412 |
|> Proof.internal_goal print_result Proof_Context.mode_schematic true "guess" |
60461 | 413 |
(SOME before_qed) after_qed |
414 |
[] [] [(Thm.empty_binding, [(Logic.mk_term goal, []), (goal, [])])] |
|
415 |
|> snd |
|
59616 | 416 |
|> Proof.refine (Method.primitive_text (fn _ => fn _ => |
60461 | 417 |
Goal.init (Thm.cterm_of ctxt thesis))) |
418 |
|> Seq.hd |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
419 |
end; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
420 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
421 |
in |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
422 |
|
60379 | 423 |
val guess = gen_guess Proof_Context.cert_var; |
424 |
val guess_cmd = gen_guess Proof_Context.read_var; |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
425 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
426 |
end; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
427 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
428 |
end; |