| author | wenzelm | 
| Thu, 30 Aug 2018 14:21:40 +0200 | |
| changeset 68853 | d36f00510e40 | 
| parent 67721 | 5348bea4accd | 
| child 71729 | 8ed68b2aeba1 | 
| permissions | -rw-r--r-- | 
| 60630 | 1  | 
(* Title: Pure/Isar/subgoal.ML  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 60623 | 3  | 
Author: Daniel Matichuk, NICTA/UNSW  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
5  | 
Tactical operations with explicit subgoal focus, based on canonical  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
6  | 
proof decomposition. The "visible" part of the text within the  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
7  | 
context is fixed, the remaining goal may be schematic.  | 
| 60623 | 8  | 
|
9  | 
Isar subgoal command for proof structure within unstructured proof  | 
|
10  | 
scripts.  | 
|
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
*)  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
signature SUBGOAL =  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
sig  | 
| 
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: 
60630 
diff
changeset
 | 
15  | 
type focus =  | 
| 
 
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: 
60630 
diff
changeset
 | 
16  | 
   {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list,
 | 
| 
 
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: 
60630 
diff
changeset
 | 
17  | 
concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list}  | 
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
18  | 
val focus_params: Proof.context -> int -> binding list option -> thm -> focus * thm  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
19  | 
val focus_params_fixed: Proof.context -> int -> binding list option -> thm -> focus * thm  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
20  | 
val focus_prems: Proof.context -> int -> binding list option -> thm -> focus * thm  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
21  | 
val focus: Proof.context -> int -> binding list option -> thm -> focus * thm  | 
| 
32210
 
a5e9d9f3e5e1
retrofit: actually handle schematic variables -- need to export into original context;
 
wenzelm 
parents: 
32200 
diff
changeset
 | 
22  | 
val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> cterm list ->  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
23  | 
int -> thm -> thm -> thm Seq.seq  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
24  | 
val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic  | 
| 
60626
 
9eefb9f39021
clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
 
wenzelm 
parents: 
60625 
diff
changeset
 | 
25  | 
val FOCUS_PARAMS_FIXED: (focus -> tactic) -> Proof.context -> int -> tactic  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
26  | 
val FOCUS_PREMS: (focus -> tactic) -> Proof.context -> int -> tactic  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
27  | 
val FOCUS: (focus -> tactic) -> Proof.context -> int -> tactic  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
28  | 
val SUBPROOF: (focus -> tactic) -> Proof.context -> int -> tactic  | 
| 60629 | 29  | 
val subgoal: Attrib.binding -> Attrib.binding option ->  | 
30  | 
bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state  | 
|
31  | 
val subgoal_cmd: Attrib.binding -> Attrib.binding option ->  | 
|
32  | 
bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state  | 
|
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
end;  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
structure Subgoal: SUBGOAL =  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
struct  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
38  | 
(* focus *)  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
39  | 
|
| 
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: 
60630 
diff
changeset
 | 
40  | 
type focus =  | 
| 
 
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: 
60630 
diff
changeset
 | 
41  | 
 {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list,
 | 
| 
 
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: 
60630 
diff
changeset
 | 
42  | 
concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list};  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
|
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
44  | 
fun gen_focus (do_prems, do_concl) ctxt i bindings raw_st =  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
let  | 
| 60367 | 46  | 
val st = raw_st  | 
| 67649 | 47  | 
|> Thm.transfer' ctxt  | 
| 60630 | 48  | 
|> Raw_Simplifier.norm_hhf_protect ctxt;  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
49  | 
val ((schematic_types, [st']), ctxt1) = Variable.importT [st] ctxt;  | 
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
50  | 
val ((params, goal), ctxt2) = Variable.focus_cterm bindings (Thm.cprem_of st' i) ctxt1;  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
51  | 
|
| 
32213
 
f1b166317ae2
added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
 
wenzelm 
parents: 
32210 
diff
changeset
 | 
52  | 
val (asms, concl) =  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
53  | 
if do_prems then (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal)  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
54  | 
else ([], goal);  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
55  | 
val text = asms @ (if do_concl then [concl] else []);  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
56  | 
|
| 60367 | 57  | 
val (inst, ctxt3) = Variable.import_inst true (map Thm.term_of text) ctxt2;  | 
| 60805 | 58  | 
val schematic_terms = map (apsnd (Thm.cterm_of ctxt3)) (#2 inst);  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
59  | 
|
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
60  | 
val schematics = (schematic_types, schematic_terms);  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
61  | 
val asms' = map (Thm.instantiate_cterm schematics) asms;  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
62  | 
val concl' = Thm.instantiate_cterm schematics concl;  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
63  | 
val (prems, context) = Assumption.add_assumes asms' ctxt3;  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
in  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
65  | 
    ({context = context, params = params, prems = prems,
 | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
66  | 
asms = asms', concl = concl', schematics = schematics}, Goal.init concl')  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
end;  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
|
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
69  | 
val focus_params = gen_focus (false, false);  | 
| 
60626
 
9eefb9f39021
clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
 
wenzelm 
parents: 
60625 
diff
changeset
 | 
70  | 
val focus_params_fixed = gen_focus (false, true);  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
71  | 
val focus_prems = gen_focus (true, false);  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
72  | 
val focus = gen_focus (true, true);  | 
| 
32213
 
f1b166317ae2
added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
 
wenzelm 
parents: 
32210 
diff
changeset
 | 
73  | 
|
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
74  | 
|
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
75  | 
(* lift and retrofit *)  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
76  | 
|
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
77  | 
(*  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
78  | 
B [?'b, ?y]  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
79  | 
----------------  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
80  | 
B ['b, y params]  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
81  | 
*)  | 
| 
32284
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
82  | 
fun lift_import idx params th ctxt =  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
83  | 
let  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
84  | 
val ((_, [th']), ctxt') = Variable.importT [th] ctxt;  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
85  | 
|
| 59586 | 86  | 
val Ts = map Thm.typ_of_cterm params;  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
87  | 
val ts = map Thm.term_of params;  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
88  | 
|
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
89  | 
val prop = Thm.full_prop_of th';  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
90  | 
val concl_vars = Term.add_vars (Logic.strip_imp_concl prop) [];  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
91  | 
val vars = rev (Term.add_vars prop []);  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
92  | 
val (ys, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) vars) ctxt';  | 
| 
32284
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
93  | 
|
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
94  | 
fun var_inst v y =  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
95  | 
let  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
96  | 
val ((x, i), T) = v;  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
97  | 
val (U, args) =  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
98  | 
if member (op =) concl_vars v then (T, [])  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
99  | 
else (Ts ---> T, ts);  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
100  | 
val u = Free (y, U);  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
101  | 
in ((Var v, list_comb (u, args)), (u, Var ((x, i + idx), U))) end;  | 
| 59616 | 102  | 
val (inst1, inst2) =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59616 
diff
changeset
 | 
103  | 
split_list (map (apply2 (apply2 (Thm.cterm_of ctxt))) (map2 var_inst vars ys));  | 
| 
32284
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
104  | 
|
| 
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: 
60630 
diff
changeset
 | 
105  | 
val th'' = Thm.instantiate ([], map (apfst (Term.dest_Var o Thm.term_of)) inst1) th';  | 
| 
32284
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
106  | 
in ((inst2, th''), ctxt'') end;  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
107  | 
|
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
108  | 
(*  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
109  | 
[x, A x]  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
110  | 
:  | 
| 67721 | 111  | 
B x \<Longrightarrow> C  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
112  | 
------------------  | 
| 67721 | 113  | 
[\<And>x. A x \<Longrightarrow> B x]  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
114  | 
:  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
115  | 
C  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
116  | 
*)  | 
| 60938 | 117  | 
fun lift_subgoals ctxt params asms th =  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
118  | 
let  | 
| 60938 | 119  | 
fun lift ct = fold_rev (Thm.all_name ctxt) params (Drule.list_implies (asms, ct));  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
120  | 
val unlift =  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
121  | 
fold (Thm.elim_implies o Thm.assume) asms o  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
122  | 
Drule.forall_elim_list (map #2 params) o Thm.assume;  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
123  | 
val subgoals = map lift (Drule.strip_imp_prems (Thm.cprop_of th));  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
124  | 
val th' = fold (Thm.elim_implies o unlift) subgoals th;  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
125  | 
in (subgoals, th') end;  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
126  | 
|
| 
32213
 
f1b166317ae2
added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
 
wenzelm 
parents: 
32210 
diff
changeset
 | 
127  | 
fun retrofit ctxt1 ctxt0 params asms i st1 st0 =  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
128  | 
let  | 
| 
32284
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
129  | 
val idx = Thm.maxidx_of st0 + 1;  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
130  | 
val ps = map #2 params;  | 
| 
32284
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
131  | 
val ((subgoal_inst, st2), ctxt2) = lift_import idx ps st1 ctxt1;  | 
| 60938 | 132  | 
val (subgoals, st3) = lift_subgoals ctxt2 params asms st2;  | 
| 
32213
 
f1b166317ae2
added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
 
wenzelm 
parents: 
32210 
diff
changeset
 | 
133  | 
val result = st3  | 
| 
 
f1b166317ae2
added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
 
wenzelm 
parents: 
32210 
diff
changeset
 | 
134  | 
|> Goal.conclude  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
135  | 
|> Drule.implies_intr_list asms  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
136  | 
|> Drule.forall_intr_list ps  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
137  | 
|> Drule.implies_intr_list subgoals  | 
| 
32284
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
138  | 
|> fold_rev (Thm.forall_intr o #1) subgoal_inst  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
139  | 
|> fold (Thm.forall_elim o #2) subgoal_inst  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
140  | 
|> Thm.adjust_maxidx_thm idx  | 
| 
 
d8ee8a956f19
retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
141  | 
|> singleton (Variable.export ctxt2 ctxt0);  | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
49845 
diff
changeset
 | 
142  | 
in  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
54883 
diff
changeset
 | 
143  | 
    Thm.bicompose (SOME ctxt0) {flatten = true, match = false, incremented = false}
 | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
49845 
diff
changeset
 | 
144  | 
(false, result, Thm.nprems_of st1) i st0  | 
| 
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
49845 
diff
changeset
 | 
145  | 
end;  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
146  | 
|
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
147  | 
|
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
148  | 
(* tacticals *)  | 
| 
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
149  | 
|
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
150  | 
fun GEN_FOCUS flags tac ctxt i st =  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
if Thm.nprems_of st < i then Seq.empty  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
else  | 
| 
60707
 
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
 
wenzelm 
parents: 
60695 
diff
changeset
 | 
153  | 
    let val (args as {context = ctxt', params, asms, ...}, st') =
 | 
| 
 
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
 
wenzelm 
parents: 
60695 
diff
changeset
 | 
154  | 
gen_focus flags (ctxt |> Variable.set_bound_focus true) i NONE st;  | 
| 32329 | 155  | 
in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end;  | 
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
156  | 
|
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
157  | 
val FOCUS_PARAMS = GEN_FOCUS (false, false);  | 
| 
60626
 
9eefb9f39021
clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
 
wenzelm 
parents: 
60625 
diff
changeset
 | 
158  | 
val FOCUS_PARAMS_FIXED = GEN_FOCUS (false, true);  | 
| 
32281
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
159  | 
val FOCUS_PREMS = GEN_FOCUS (true, false);  | 
| 
 
750101435f60
focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
 
wenzelm 
parents: 
32213 
diff
changeset
 | 
160  | 
val FOCUS = GEN_FOCUS (true, true);  | 
| 
32213
 
f1b166317ae2
added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
 
wenzelm 
parents: 
32210 
diff
changeset
 | 
161  | 
|
| 49845 | 162  | 
fun SUBPROOF tac ctxt = FOCUS (Seq.map (Goal.check_finished ctxt) oo tac) ctxt;  | 
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
|
| 60623 | 164  | 
|
165  | 
(* Isar subgoal command *)  | 
|
166  | 
||
167  | 
local  | 
|
168  | 
||
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
169  | 
fun param_bindings ctxt (param_suffix, raw_param_specs) st =  | 
| 
60628
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
170  | 
let  | 
| 
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
171  | 
val _ = if Thm.no_prems st then error "No subgoals!" else ();  | 
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
172  | 
val subgoal = #1 (Logic.dest_implies (Thm.prop_of st));  | 
| 60629 | 173  | 
val subgoal_params =  | 
174  | 
map (apfst (Name.internal o Name.clean)) (Term.strip_all_vars subgoal)  | 
|
175  | 
|> Term.variant_frees subgoal |> map #1;  | 
|
| 
60628
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
176  | 
|
| 60629 | 177  | 
val n = length subgoal_params;  | 
178  | 
val m = length raw_param_specs;  | 
|
| 
60628
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
179  | 
val _ =  | 
| 60629 | 180  | 
m <= n orelse  | 
181  | 
        error ("Excessive subgoal parameter specification" ^
 | 
|
182  | 
Position.here_list (map snd (drop n raw_param_specs)));  | 
|
| 
60628
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
183  | 
|
| 60629 | 184  | 
val param_specs =  | 
185  | 
raw_param_specs |> map  | 
|
186  | 
(fn (NONE, _) => NONE  | 
|
187  | 
| (SOME x, pos) =>  | 
|
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
188  | 
let  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
189  | 
val b = #1 (#1 (Proof_Context.cert_var (Binding.make (x, pos), NONE, NoSyn) ctxt));  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
190  | 
val _ = Variable.check_name b;  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
191  | 
in SOME b end)  | 
| 60629 | 192  | 
|> param_suffix ? append (replicate (n - m) NONE);  | 
193  | 
||
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
194  | 
fun bindings (SOME x :: xs) (_ :: ys) = x :: bindings xs ys  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
195  | 
| bindings (NONE :: xs) (y :: ys) = Binding.name y :: bindings xs ys  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
196  | 
| bindings _ ys = map Binding.name ys;  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
197  | 
in bindings param_specs subgoal_params end;  | 
| 
60628
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
198  | 
|
| 
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
199  | 
fun gen_subgoal prep_atts raw_result_binding raw_prems_binding param_specs state =  | 
| 60623 | 200  | 
let  | 
201  | 
val _ = Proof.assert_backward state;  | 
|
| 
60627
 
5d13babbb3f6
split multi-goals as usual (outermost Pure.conjunction only);
 
wenzelm 
parents: 
60626 
diff
changeset
 | 
202  | 
|
| 
 
5d13babbb3f6
split multi-goals as usual (outermost Pure.conjunction only);
 
wenzelm 
parents: 
60626 
diff
changeset
 | 
203  | 
val state1 = state |> Proof.refine_insert [];  | 
| 
 
5d13babbb3f6
split multi-goals as usual (outermost Pure.conjunction only);
 
wenzelm 
parents: 
60626 
diff
changeset
 | 
204  | 
    val {context = ctxt, facts = facts, goal = st} = Proof.raw_goal state1;
 | 
| 60623 | 205  | 
|
206  | 
val result_binding = apsnd (map (prep_atts ctxt)) raw_result_binding;  | 
|
207  | 
val (prems_binding, do_prems) =  | 
|
208  | 
(case raw_prems_binding of  | 
|
209  | 
SOME (b, raw_atts) => ((b, map (prep_atts ctxt) raw_atts), true)  | 
|
| 63352 | 210  | 
| NONE => (Binding.empty_atts, false));  | 
| 60623 | 211  | 
|
| 
60628
 
5ff15d0dd7fa
subgoal parameters are internal by default and named by user;
 
wenzelm 
parents: 
60627 
diff
changeset
 | 
212  | 
val (subgoal_focus, _) =  | 
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
213  | 
(if do_prems then focus else focus_params_fixed) ctxt  | 
| 
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
214  | 
1 (SOME (param_bindings ctxt param_specs st)) st;  | 
| 60623 | 215  | 
|
216  | 
fun after_qed (ctxt'', [[result]]) =  | 
|
217  | 
Proof.end_block #> (fn state' =>  | 
|
218  | 
let  | 
|
219  | 
val ctxt' = Proof.context_of state';  | 
|
220  | 
val results' =  | 
|
221  | 
Proof_Context.export ctxt'' ctxt' (Conjunction.elim_conjunctions result);  | 
|
222  | 
in  | 
|
223  | 
state'  | 
|
224  | 
|> Proof.refine_primitive (fn _ => fn _ =>  | 
|
225  | 
retrofit ctxt'' ctxt' (#params subgoal_focus) (#asms subgoal_focus) 1  | 
|
226  | 
(Goal.protect 0 result) st  | 
|
227  | 
|> Seq.hd)  | 
|
228  | 
|> Proof.map_context  | 
|
229  | 
(#2 o Proof_Context.note_thmss "" [(result_binding, [(results', [])])])  | 
|
| 60625 | 230  | 
end)  | 
231  | 
#> Proof.reset_facts  | 
|
232  | 
#> Proof.enter_backward;  | 
|
| 60623 | 233  | 
in  | 
| 
60627
 
5d13babbb3f6
split multi-goals as usual (outermost Pure.conjunction only);
 
wenzelm 
parents: 
60626 
diff
changeset
 | 
234  | 
state1  | 
| 60623 | 235  | 
|> Proof.enter_forward  | 
236  | 
|> Proof.using_facts []  | 
|
237  | 
|> Proof.begin_block  | 
|
238  | 
|> Proof.map_context (fn _ =>  | 
|
239  | 
#context subgoal_focus  | 
|
240  | 
|> Proof_Context.note_thmss "" [(prems_binding, [(#prems subgoal_focus, [])])] |> #2)  | 
|
241  | 
|> Proof.internal_goal (K (K ())) (Proof_Context.get_mode ctxt) true "subgoal"  | 
|
| 63352 | 242  | 
NONE after_qed [] [] [(Binding.empty_atts, [(Thm.term_of (#concl subgoal_focus), [])])] |> #2  | 
| 60623 | 243  | 
|> Proof.using_facts facts  | 
244  | 
|> pair subgoal_focus  | 
|
245  | 
end;  | 
|
246  | 
||
247  | 
in  | 
|
248  | 
||
| 
60626
 
9eefb9f39021
clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
 
wenzelm 
parents: 
60625 
diff
changeset
 | 
249  | 
val subgoal = gen_subgoal Attrib.attribute;  | 
| 
 
9eefb9f39021
clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
 
wenzelm 
parents: 
60625 
diff
changeset
 | 
250  | 
val subgoal_cmd = gen_subgoal Attrib.attribute_cmd;  | 
| 60623 | 251  | 
|
252  | 
end;  | 
|
253  | 
||
| 
20210
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
254  | 
end;  | 
| 
 
8fe4ae4360eb
Tactical operations depending on local subgoal structure.
 
wenzelm 
parents:  
diff
changeset
 | 
255  | 
|
| 
32200
 
2bd8ab91a426
advanced retrofit, which allows new subgoals and variables;
 
wenzelm 
parents: 
31794 
diff
changeset
 | 
256  | 
val SUBPROOF = Subgoal.SUBPROOF;  |