| author | wenzelm | 
| Sun, 23 Jul 2023 21:04:33 +0200 | |
| changeset 78445 | d70dc78c0d4e | 
| parent 70521 | 9ddd66d53130 | 
| permissions | -rw-r--r-- | 
| 70520 | 1  | 
(* Title: Pure/context_tactic.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Tactics with proof context / cases -- as basis for Isar proof methods.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
infix 1 CONTEXT_THEN_ALL_NEW;  | 
|
8  | 
||
9  | 
signature BASIC_CONTEXT_TACTIC =  | 
|
10  | 
sig  | 
|
11  | 
type context_state = Proof.context * thm  | 
|
12  | 
type context_tactic = context_state -> context_state Seq.result Seq.seq  | 
|
13  | 
val TACTIC_CONTEXT: Proof.context -> thm Seq.seq -> context_state Seq.result Seq.seq  | 
|
14  | 
val CONTEXT_TACTIC: tactic -> context_tactic  | 
|
15  | 
val NO_CONTEXT_TACTIC: Proof.context -> context_tactic -> tactic  | 
|
16  | 
val CONTEXT_CASES: Rule_Cases.cases -> tactic -> context_tactic  | 
|
17  | 
val CONTEXT_SUBGOAL: (term * int -> context_tactic) -> int -> context_tactic  | 
|
18  | 
val CONTEXT_THEN_ALL_NEW: (int -> context_tactic) * (int -> tactic) -> int -> context_tactic  | 
|
| 
70521
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
19  | 
val SUBPROOFS: context_tactic -> context_tactic  | 
| 70520 | 20  | 
end;  | 
21  | 
||
22  | 
signature CONTEXT_TACTIC =  | 
|
23  | 
sig  | 
|
24  | 
include BASIC_CONTEXT_TACTIC  | 
|
25  | 
end;  | 
|
26  | 
||
27  | 
structure Context_Tactic: CONTEXT_TACTIC =  | 
|
28  | 
struct  | 
|
29  | 
||
| 
70521
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
30  | 
(* type context_tactic *)  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
31  | 
|
| 70520 | 32  | 
type context_state = Proof.context * thm;  | 
33  | 
type context_tactic = context_state -> context_state Seq.result Seq.seq;  | 
|
34  | 
||
35  | 
fun TACTIC_CONTEXT ctxt : thm Seq.seq -> context_state Seq.result Seq.seq =  | 
|
36  | 
Seq.map (Seq.Result o pair ctxt);  | 
|
37  | 
||
38  | 
fun CONTEXT_TACTIC tac : context_tactic =  | 
|
39  | 
fn (ctxt, st) => TACTIC_CONTEXT ctxt (tac st);  | 
|
40  | 
||
41  | 
fun NO_CONTEXT_TACTIC ctxt (tac: context_tactic) st =  | 
|
42  | 
tac (ctxt, st) |> Seq.filter_results |> Seq.map snd;  | 
|
43  | 
||
44  | 
fun CONTEXT_CASES cases tac : context_tactic =  | 
|
45  | 
fn (ctxt, st) => TACTIC_CONTEXT (Proof_Context.update_cases cases ctxt) (tac st);  | 
|
46  | 
||
47  | 
fun CONTEXT_SUBGOAL tac i : context_tactic =  | 
|
48  | 
fn (ctxt, st) =>  | 
|
49  | 
(case try Logic.nth_prem (i, Thm.prop_of st) of  | 
|
50  | 
SOME goal => tac (goal, i) (ctxt, st)  | 
|
51  | 
| NONE => Seq.empty);  | 
|
52  | 
||
53  | 
fun (tac1 CONTEXT_THEN_ALL_NEW tac2) i : context_tactic =  | 
|
54  | 
fn (ctxt, st) =>  | 
|
55  | 
(ctxt, st) |> tac1 i |> Seq.maps_results (fn (ctxt', st') =>  | 
|
56  | 
TACTIC_CONTEXT ctxt' ((Seq.INTERVAL tac2 i (i + Thm.nprems_of st' - Thm.nprems_of st)) st'));  | 
|
57  | 
||
| 
70521
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
58  | 
|
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
59  | 
(* subproofs with closed derivation *)  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
60  | 
|
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
61  | 
fun SUBPROOFS tac : context_tactic =  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
62  | 
let  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
63  | 
fun apply (g :: gs) (SOME (Seq.Result (results, ctxt))) =  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
64  | 
(case Seq.pull (tac (ctxt, Goal.init g)) of  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
65  | 
SOME (Seq.Result (ctxt', st'), _) =>  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
66  | 
apply gs (SOME (Seq.Result (st' :: results, ctxt')))  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
67  | 
| SOME (Seq.Error msg, _) => SOME (Seq.Error msg)  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
68  | 
| NONE => NONE)  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
69  | 
| apply _ x = x;  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
70  | 
in  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
71  | 
fn (ctxt, st) =>  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
72  | 
(case Par_Tactical.strip_goals st of  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
73  | 
SOME goals =>  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
74  | 
(case apply goals (SOME (Seq.Result ([], ctxt))) of  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
75  | 
SOME (Seq.Result (results, ctxt')) =>  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
76  | 
              TACTIC_CONTEXT ctxt' (Par_Tactical.retrofit_tac {close = true} results st)
 | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
77  | 
| SOME (Seq.Error msg) => Seq.single (Seq.Error msg)  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
78  | 
| NONE => Seq.empty)  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
79  | 
| NONE => Seq.DETERM tac (ctxt, st))  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
80  | 
end;  | 
| 
 
9ddd66d53130
added SUBPROOFS / "subproofs" method combinator, for more compact proofterms;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
81  | 
|
| 70520 | 82  | 
end;  | 
83  | 
||
84  | 
structure Basic_Context_Tactic: BASIC_CONTEXT_TACTIC = Context_Tactic;  | 
|
85  | 
open Basic_Context_Tactic;  |