author | haftmann |
Mon, 08 Oct 2007 08:04:28 +0200 | |
changeset 24901 | d3cbf79769b9 |
parent 22568 | ed7aa5a350ef |
child 29606 | fedb8be05f24 |
permissions | -rw-r--r-- |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/subgoal.ML |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
3 |
Author: Makarius |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
4 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
5 |
Tactical operations depending on local subgoal structure. |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
6 |
*) |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
7 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
8 |
signature BASIC_SUBGOAL = |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
9 |
sig |
20231 | 10 |
val SUBPROOF: |
20219
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
11 |
({context: Proof.context, schematics: ctyp list * cterm list, |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
12 |
params: cterm list, asms: cterm list, concl: cterm, |
20231 | 13 |
prems: thm list} -> tactic) -> Proof.context -> int -> tactic |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
14 |
end |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
15 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
16 |
signature SUBGOAL = |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
17 |
sig |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
18 |
include BASIC_SUBGOAL |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
19 |
val focus: Proof.context -> int -> thm -> |
20219
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
20 |
{context: Proof.context, schematics: ctyp list * cterm list, |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
21 |
params: cterm list, asms: cterm list, concl: cterm, prems: thm list} * thm |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
22 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
23 |
end; |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
24 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
25 |
structure Subgoal: SUBGOAL = |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
26 |
struct |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
27 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
28 |
(* canonical proof decomposition -- similar to fix/assume/show *) |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
29 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
30 |
fun focus ctxt i st = |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
31 |
let |
21605 | 32 |
val ((schematics, [st']), ctxt') = |
22568
ed7aa5a350ef
renamed Variable.import to import_thms (avoid clash with Alice keywords);
wenzelm
parents:
21605
diff
changeset
|
33 |
Variable.import_thms true [MetaSimplifier.norm_hhf_protect st] ctxt; |
20301 | 34 |
val ((params, goal), ctxt'') = Variable.focus_subgoal i st' ctxt'; |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
35 |
val asms = Drule.strip_imp_prems goal; |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
36 |
val concl = Drule.strip_imp_concl goal; |
20231 | 37 |
val (prems, context) = Assumption.add_assumes asms ctxt''; |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
38 |
in |
20219
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
39 |
({context = context, schematics = schematics, params = params, |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
40 |
asms = asms, concl = concl, prems = prems}, Goal.init concl) |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
41 |
end; |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
42 |
|
20231 | 43 |
fun SUBPROOF tac ctxt i st = |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
44 |
if Thm.nprems_of st < i then Seq.empty |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
45 |
else |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
46 |
let |
20219
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
47 |
val (args as {context, params, ...}, st') = focus ctxt i st; |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
48 |
fun export_closed th = |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
49 |
let |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
50 |
val (th' :: params') = ProofContext.export context ctxt (th :: map Drule.mk_term params); |
20579 | 51 |
val vs = map (Thm.dest_arg o Drule.strip_imp_concl o Thm.cprop_of) params'; |
20219
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
52 |
in Drule.forall_intr_list vs th' end; |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
53 |
fun retrofit th = Thm.compose_no_flatten false (th, 0) i; |
3bff4719f3d6
focus: result record includes (fixed) schematic variables;
wenzelm
parents:
20210
diff
changeset
|
54 |
in Seq.lifts retrofit (Seq.map (Goal.finish #> export_closed) (tac args st')) st end |
20210
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
55 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
56 |
end; |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
57 |
|
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
58 |
structure BasicSubgoal: BASIC_SUBGOAL = Subgoal; |
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff
changeset
|
59 |
open BasicSubgoal; |