author | haftmann |
Sat, 20 Aug 2011 01:21:22 +0200 | |
changeset 44323 | 4b5b430eb00e |
parent 43324 | 2b47822868e4 |
child 45327 | 4a027cc86f1a |
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 |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
4 |
The 'obtain' and 'guess' language elements -- generalized existence at |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
5 |
the level of proof texts: 'obtain' involves a proof that certain |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
6 |
fixes/assumes may be introduced into the present context; 'guess' is |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
7 |
similar, but derives these elements from the course of reasoning! |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
8 |
|
9468 | 9 |
<chain_facts> |
18870 | 10 |
obtain x where "A x" <proof> == |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
11 |
|
18870 | 12 |
have "!!thesis. (!!x. A x ==> thesis) ==> thesis" |
12970 | 13 |
proof succeed |
9468 | 14 |
fix thesis |
18870 | 15 |
assume that [intro?]: "!!x. A x ==> thesis" |
16 |
<chain_facts> |
|
17 |
show thesis |
|
18 |
apply (insert that) |
|
19 |
<proof> |
|
12970 | 20 |
qed |
18870 | 21 |
fix x assm <<obtain_export>> "A x" |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
22 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
23 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
24 |
<chain_facts> |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
25 |
guess x <proof body> <proof end> == |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
26 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
27 |
{ |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
28 |
fix thesis |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
29 |
<chain_facts> have "PROP ?guess" |
18870 | 30 |
apply magic -- {* turns goal into "thesis ==> #thesis" *} |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
31 |
<proof body> |
18870 | 32 |
apply_end magic -- {* turns final "(!!x. P x ==> thesis) ==> #thesis" into |
33 |
"#((!!x. A x ==> thesis) ==> thesis)" which is a finished goal state *} |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
34 |
<proof end> |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
35 |
} |
18870 | 36 |
fix x assm <<obtain_export>> "A x" |
8094 | 37 |
*) |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
38 |
|
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
39 |
signature OBTAIN = |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
40 |
sig |
21229 | 41 |
val thatN: string |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
42 |
val obtain: string -> (binding * typ option * mixfix) list -> |
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
43 |
(Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state |
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
44 |
val obtain_cmd: string -> (binding * string option * mixfix) list -> |
30211 | 45 |
(Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state |
20308 | 46 |
val result: (Proof.context -> tactic) -> thm list -> Proof.context -> |
32199 | 47 |
((string * cterm) list * thm list) * Proof.context |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
48 |
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
|
49 |
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
|
50 |
end; |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
51 |
|
10379
93630e0c5ae9
improved handling of "that": insert into goal, only declare as Pure "intro";
wenzelm
parents:
9481
diff
changeset
|
52 |
structure Obtain: OBTAIN = |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
53 |
struct |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
54 |
|
18670 | 55 |
(** obtain_export **) |
56 |
||
18870 | 57 |
(* |
18897 | 58 |
[x, A x] |
59 |
: |
|
60 |
B |
|
61 |
-------- |
|
62 |
B |
|
18870 | 63 |
*) |
21686 | 64 |
fun eliminate_term ctxt xs tm = |
65 |
let |
|
66 |
val vs = map (dest_Free o Thm.term_of) xs; |
|
67 |
val bads = Term.fold_aterms (fn t as Free v => |
|
68 |
if member (op =) vs v then insert (op aconv) t else I | _ => I) tm []; |
|
69 |
val _ = null bads orelse |
|
70 |
error ("Result contains obtained parameters: " ^ |
|
24920 | 71 |
space_implode " " (map (Syntax.string_of_term ctxt) bads)); |
21686 | 72 |
in tm end; |
73 |
||
74 |
fun eliminate fix_ctxt rule xs As thm = |
|
9468 | 75 |
let |
42360 | 76 |
val thy = Proof_Context.theory_of fix_ctxt; |
9468 | 77 |
|
21686 | 78 |
val _ = eliminate_term fix_ctxt xs (Thm.full_prop_of thm); |
35625 | 79 |
val _ = Object_Logic.is_judgment thy (Thm.concl_of thm) orelse |
20308 | 80 |
error "Conclusion in obtained context must be object-logic judgment"; |
81 |
||
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30763
diff
changeset
|
82 |
val ((_, [thm']), ctxt') = Variable.import true [thm] fix_ctxt; |
20308 | 83 |
val prems = Drule.strip_imp_prems (#prop (Thm.crep_thm thm')); |
9468 | 84 |
in |
20308 | 85 |
((Drule.implies_elim_list thm' (map Thm.assume prems) |
86 |
|> Drule.implies_intr_list (map Drule.norm_hhf_cterm As) |
|
87 |
|> Drule.forall_intr_list xs) |
|
88 |
COMP rule) |
|
89 |
|> Drule.implies_intr_list prems |
|
90 |
|> singleton (Variable.export ctxt' fix_ctxt) |
|
9468 | 91 |
end; |
92 |
||
21686 | 93 |
fun obtain_export ctxt rule xs _ As = |
94 |
(eliminate ctxt rule xs As, eliminate_term ctxt xs); |
|
95 |
||
9468 | 96 |
|
97 |
||
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
98 |
(** obtain **) |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
99 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
100 |
fun bind_judgment ctxt name = |
18670 | 101 |
let |
42501
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
102 |
val thy = Proof_Context.theory_of ctxt; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
103 |
val ([x], ctxt') = Proof_Context.add_fixes [(Binding.name name, NONE, NoSyn)] ctxt; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
104 |
val (t as _ $ Free v) = Object_Logic.fixed_judgment thy x; |
20308 | 105 |
in ((v, t), ctxt') end; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
106 |
|
18897 | 107 |
val thatN = "that"; |
108 |
||
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
109 |
local |
8094 | 110 |
|
18897 | 111 |
fun gen_obtain prep_att prep_vars prep_propp |
112 |
name raw_vars raw_asms int state = |
|
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
113 |
let |
9468 | 114 |
val _ = Proof.assert_forward_or_chain state; |
20308 | 115 |
val thy = Proof.theory_of state; |
116 |
val cert = Thm.cterm_of thy; |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
117 |
val ctxt = Proof.context_of state; |
17357 | 118 |
val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else []; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
119 |
|
8543 | 120 |
(*obtain vars*) |
19844 | 121 |
val (vars, vars_ctxt) = prep_vars raw_vars ctxt; |
42490
3633ecaaf3ef
more formal treatment of parameters, avoiding slightly odd Variable.intern_fixed;
wenzelm
parents:
42488
diff
changeset
|
122 |
val (xs', fix_ctxt) = vars_ctxt |> Proof_Context.add_fixes vars; |
42494 | 123 |
val xs = map (Variable.check_name o #1) vars; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
124 |
|
8543 | 125 |
(*obtain asms*) |
11890 | 126 |
val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms); |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19300
diff
changeset
|
127 |
val asm_props = maps (map fst) proppss; |
18678 | 128 |
val asms = map fst (Attrib.map_specs (prep_att thy) raw_asms) ~~ proppss; |
10464 | 129 |
|
42490
3633ecaaf3ef
more formal treatment of parameters, avoiding slightly odd Variable.intern_fixed;
wenzelm
parents:
42488
diff
changeset
|
130 |
(*obtain parms*) |
3633ecaaf3ef
more formal treatment of parameters, avoiding slightly odd Variable.intern_fixed;
wenzelm
parents:
42488
diff
changeset
|
131 |
val (Ts, parms_ctxt) = fold_map Proof_Context.inferred_param xs' asms_ctxt; |
3633ecaaf3ef
more formal treatment of parameters, avoiding slightly odd Variable.intern_fixed;
wenzelm
parents:
42488
diff
changeset
|
132 |
val parms = xs' ~~ Ts; |
3633ecaaf3ef
more formal treatment of parameters, avoiding slightly odd Variable.intern_fixed;
wenzelm
parents:
42488
diff
changeset
|
133 |
val _ = Variable.warn_extra_tfrees fix_ctxt parms_ctxt; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
134 |
|
12970 | 135 |
(*obtain statements*) |
43324
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents:
42501
diff
changeset
|
136 |
val thesisN = singleton (Name.variant_list xs) Auto_Bind.thesisN; |
20308 | 137 |
val (thesis_var, thesis) = #1 (bind_judgment fix_ctxt thesisN); |
9468 | 138 |
|
18897 | 139 |
val that_name = if name = "" then thatN else name; |
10582 | 140 |
val that_prop = |
20308 | 141 |
Term.list_all_free (parms, Logic.list_implies (asm_props, thesis)) |
142 |
|> Library.curry Logic.list_rename_params xs; |
|
12970 | 143 |
val obtain_prop = |
33386 | 144 |
Logic.list_rename_params ([Auto_Bind.thesisN], |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
145 |
Term.list_all_free ([thesis_var], Logic.mk_implies (that_prop, thesis))); |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
146 |
|
18124 | 147 |
fun after_qed _ = |
17357 | 148 |
Proof.local_qed (NONE, false) |
29383 | 149 |
#> `Proof.the_fact #-> (fn rule => |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
150 |
Proof.fix vars |
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
151 |
#> Proof.assm (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms); |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
152 |
in |
8094 | 153 |
state |
9468 | 154 |
|> Proof.enter_forward |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
155 |
|> Proof.have NONE (K I) [(Thm.empty_binding, [(obtain_prop, [])])] int |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
156 |
|> Proof.proof (SOME Method.succeed_text) |> Seq.hd |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
157 |
|> Proof.fix [(Binding.name thesisN, NONE, NoSyn)] |
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
158 |
|> Proof.assume |
33369 | 159 |
[((Binding.name that_name, [Context_Rules.intro_query NONE]), [(that_prop, [])])] |
16842 | 160 |
|> `Proof.the_facts |
17357 | 161 |
||> Proof.chain_facts chain_facts |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
162 |
||> Proof.show NONE after_qed [(Thm.empty_binding, [(thesis, [])])] false |
18907 | 163 |
|-> Proof.refine_insert |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
164 |
end; |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
165 |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
166 |
in |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
167 |
|
42360 | 168 |
val obtain = gen_obtain (K I) Proof_Context.cert_vars Proof_Context.cert_propp; |
169 |
val obtain_cmd = gen_obtain Attrib.attribute Proof_Context.read_vars Proof_Context.read_propp; |
|
8094 | 170 |
|
171 |
end; |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
172 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
173 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
174 |
|
20308 | 175 |
(** tactical result **) |
176 |
||
177 |
fun check_result ctxt thesis th = |
|
178 |
(case Thm.prems_of th of |
|
179 |
[prem] => |
|
180 |
if Thm.concl_of th aconv thesis andalso |
|
181 |
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
|
182 |
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
|
183 |
| [] => 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
|
184 |
| _ => error ("Guess split into several cases:\n" ^ Display.string_of_thm ctxt th)); |
20308 | 185 |
|
186 |
fun result tac facts ctxt = |
|
187 |
let |
|
42360 | 188 |
val thy = Proof_Context.theory_of ctxt; |
20308 | 189 |
val cert = Thm.cterm_of thy; |
190 |
||
33386 | 191 |
val ((thesis_var, thesis), thesis_ctxt) = bind_judgment ctxt Auto_Bind.thesisN; |
20308 | 192 |
val rule = |
193 |
(case SINGLE (Method.insert_tac facts 1 THEN tac thesis_ctxt) (Goal.init (cert thesis)) of |
|
194 |
NONE => raise THM ("Obtain.result: tactic failed", 0, facts) |
|
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39134
diff
changeset
|
195 |
| SOME th => check_result ctxt thesis (Raw_Simplifier.norm_hhf (Goal.conclude th))); |
20308 | 196 |
|
197 |
val closed_rule = Thm.forall_intr (cert (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
|
198 |
val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt; |
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35625
diff
changeset
|
199 |
val obtain_rule = Thm.forall_elim (cert (Logic.varify_global (Free thesis_var))) rule'; |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
200 |
val ((params, stmt), fix_ctxt) = Variable.focus_cterm (Thm.cprem_of obtain_rule 1) ctxt'; |
20308 | 201 |
val (prems, ctxt'') = |
32199 | 202 |
Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params)) |
20308 | 203 |
(Drule.strip_imp_prems stmt) fix_ctxt; |
204 |
in ((params, prems), ctxt'') end; |
|
205 |
||
206 |
||
207 |
||
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
208 |
(** guess **) |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
209 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
210 |
local |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
211 |
|
20308 | 212 |
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
|
213 |
let |
42360 | 214 |
val thy = Proof_Context.theory_of ctxt; |
19978 | 215 |
val certT = Thm.ctyp_of thy; |
216 |
val cert = Thm.cterm_of thy; |
|
39134
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
wenzelm
parents:
38875
diff
changeset
|
217 |
val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); |
17891 | 218 |
|
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
|
219 |
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
|
220 |
|
19978 | 221 |
val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1; |
19779 | 222 |
val rule = Thm.incr_indexes (maxidx + 1) raw_rule; |
223 |
||
33368 | 224 |
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
|
225 |
val m = length vars; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
226 |
val n = length params; |
19779 | 227 |
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
|
228 |
|
19779 | 229 |
fun unify ((x, T), (y, U)) (tyenv, max) = Sign.typ_unify thy (T, U) (tyenv, max) |
230 |
handle Type.TUNIFY => |
|
231 |
err ("Failed to unify variable " ^ |
|
232 |
string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^ |
|
42284 | 233 |
string_of_term (Syntax_Trans.mark_boundT (y, Envir.norm_type tyenv U)) ^ " in") rule; |
33957 | 234 |
val (tyenv, _) = fold unify (map #1 vars ~~ take m params) |
19779 | 235 |
(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
|
236 |
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
|
237 |
|
19978 | 238 |
val xs = map (apsnd norm_type o fst) vars; |
33957 | 239 |
val ys = map (apsnd norm_type) (drop m params); |
20085
c5d60752587f
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
20004
diff
changeset
|
240 |
val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys; |
19978 | 241 |
val terms = map (Drule.mk_term o cert o Free) (xs @ ys'); |
19779 | 242 |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
243 |
val instT = |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
244 |
fold (Term.add_tvarsT o #2) params [] |
19978 | 245 |
|> map (TVar #> (fn T => (certT T, certT (norm_type T)))); |
20308 | 246 |
val closed_rule = rule |
247 |
|> Thm.forall_intr (cert (Free thesis_var)) |
|
248 |
|> Thm.instantiate (instT, []); |
|
17891 | 249 |
|
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30763
diff
changeset
|
250 |
val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt; |
19978 | 251 |
val vars' = |
252 |
map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~ |
|
253 |
(map snd vars @ replicate (length ys) NoSyn); |
|
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35625
diff
changeset
|
254 |
val rule'' = Thm.forall_elim (cert (Logic.varify_global (Free thesis_var))) rule'; |
19978 | 255 |
in ((vars', rule''), ctxt') end; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
256 |
|
28080
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
24920
diff
changeset
|
257 |
fun inferred_type (binding, _, mx) ctxt = |
4723eb2456ce
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
24920
diff
changeset
|
258 |
let |
42494 | 259 |
val x = Variable.check_name binding; |
42360 | 260 |
val (T, ctxt') = Proof_Context.inferred_param x ctxt |
19779 | 261 |
in ((x, T, mx), ctxt') end; |
262 |
||
20004
e6d3f2b031e6
guess: proper context for polymorphic parameters;
wenzelm
parents:
19978
diff
changeset
|
263 |
fun polymorphic ctxt vars = |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19844
diff
changeset
|
264 |
let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars)) |
19779 | 265 |
in map2 (fn (x, _, mx) => fn T => ((x, T), mx)) vars Ts end; |
18693 | 266 |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
267 |
fun gen_guess prep_vars raw_vars int state = |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
268 |
let |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
269 |
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
|
270 |
val thy = Proof.theory_of state; |
20308 | 271 |
val cert = Thm.cterm_of thy; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
272 |
val ctxt = Proof.context_of state; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
273 |
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
|
274 |
|
33386 | 275 |
val (thesis_var, thesis) = #1 (bind_judgment ctxt Auto_Bind.thesisN); |
20004
e6d3f2b031e6
guess: proper context for polymorphic parameters;
wenzelm
parents:
19978
diff
changeset
|
276 |
val vars = ctxt |> prep_vars raw_vars |-> fold_map inferred_type |> fst |> polymorphic ctxt; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
277 |
|
19978 | 278 |
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
|
279 |
let |
19978 | 280 |
val ((parms, rule), ctxt') = |
20308 | 281 |
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
|
282 |
val (xs, _) = Variable.add_fixes (map (#1 o #1) parms) ctxt'; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
283 |
val ps = xs ~~ map (#2 o #1) parms; |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
284 |
val ts = map Free ps; |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
285 |
val asms = |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
286 |
Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule)) |
19585 | 287 |
|> map (fn asm => (Term.betapplys (Term.list_abs (ps, asm), ts), [])); |
19779 | 288 |
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
|
289 |
in |
19978 | 290 |
state' |
291 |
|> Proof.map_context (K ctxt') |
|
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
292 |
|> 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
|
293 |
|> `Proof.context_of |-> (fn fix_ctxt => Proof.assm |
30211 | 294 |
(obtain_export fix_ctxt rule (map cert ts)) [(Thm.empty_binding, asms)]) |
33386 | 295 |
|> Proof.bind_terms Auto_Bind.no_facts |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
296 |
end; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
297 |
|
19779 | 298 |
val goal = Var (("guess", 0), propT); |
299 |
fun print_result ctxt' (k, [(s, [_, th])]) = |
|
33389 | 300 |
Proof_Display.print_results int ctxt' (k, [(s, [th])]); |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39134
diff
changeset
|
301 |
val before_qed = SOME (Method.primitive_text (Goal.conclude #> Raw_Simplifier.norm_hhf #> |
20308 | 302 |
(fn th => Goal.protect (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th)))); |
19779 | 303 |
fun after_qed [[_, res]] = |
29383 | 304 |
Proof.end_block #> guess_context (check_result ctxt thesis res); |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
305 |
in |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
306 |
state |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
307 |
|> Proof.enter_forward |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
308 |
|> Proof.begin_block |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
35845
diff
changeset
|
309 |
|> 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
|
310 |
|> Proof.chain_facts chain_facts |
19779 | 311 |
|> Proof.local_goal print_result (K I) (apsnd (rpair I)) |
30211 | 312 |
"guess" before_qed after_qed [(Thm.empty_binding, [Logic.mk_term goal, goal])] |
20308 | 313 |
|> Proof.refine (Method.primitive_text (K (Goal.init (cert thesis)))) |> Seq.hd |
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
314 |
end; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
315 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
316 |
in |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
317 |
|
42360 | 318 |
val guess = gen_guess Proof_Context.cert_vars; |
319 |
val guess_cmd = gen_guess Proof_Context.read_vars; |
|
17858
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
320 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
321 |
end; |
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
322 |
|
bc4db8cfd92f
added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents:
17357
diff
changeset
|
323 |
end; |