author | wenzelm |
Thu, 22 Dec 2005 00:29:09 +0100 | |
changeset 18477 | bf2a02c82a55 |
parent 18375 | 99deeed095ae |
child 18581 | 6538fdcc98dc |
permissions | -rw-r--r-- |
8364 | 1 |
(* Title: Pure/Isar/rule_cases.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
5 |
Annotations and local contexts of rules. |
8364 | 6 |
*) |
7 |
||
18188 | 8 |
infix 1 THEN_ALL_NEW_CASES; |
9 |
||
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
10 |
signature BASIC_RULE_CASES = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
11 |
sig |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
12 |
type cases |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
13 |
type cases_tactic |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
14 |
val CASES: cases -> tactic -> cases_tactic |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
15 |
val NO_CASES: tactic -> cases_tactic |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
16 |
val SUBGOAL_CASES: ((term * int) -> cases_tactic) -> int -> cases_tactic |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
17 |
val THEN_ALL_NEW_CASES: (int -> cases_tactic) * (int -> tactic) -> int -> cases_tactic |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
18 |
end |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
19 |
|
8364 | 20 |
signature RULE_CASES = |
21 |
sig |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
22 |
include BASIC_RULE_CASES |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
23 |
type T |
18237 | 24 |
val consume: thm list -> thm list -> ('a * int) * thm -> |
25 |
(('a * (int * thm list)) * thm) Seq.seq |
|
18477 | 26 |
val add_consumes: int -> thm -> thm |
10530 | 27 |
val consumes: int -> 'a attribute |
28 |
val consumes_default: int -> 'a attribute |
|
8364 | 29 |
val name: string list -> thm -> thm |
8427 | 30 |
val case_names: string list -> 'a attribute |
18237 | 31 |
val case_conclusion: string * string list -> 'a attribute |
12761 | 32 |
val save: thm -> thm -> thm |
18237 | 33 |
val get: thm -> (string * string list) list * int |
17861 | 34 |
val strip_params: term -> (string * typ) list |
18237 | 35 |
val make: bool -> term option -> theory * term -> (string * string list) list -> cases |
18477 | 36 |
val simple: theory * term -> string -> cases |
8427 | 37 |
val rename_params: string list list -> thm -> thm |
8364 | 38 |
val params: string list list -> 'a attribute |
18477 | 39 |
val mutual_rule: thm list -> (int list * thm) option |
8364 | 40 |
end; |
41 |
||
42 |
structure RuleCases: RULE_CASES = |
|
43 |
struct |
|
44 |
||
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
45 |
(** tactics with cases **) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
46 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
47 |
type T = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
48 |
{fixes: (string * typ) list, |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
49 |
assumes: (string * term list) list, |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
50 |
binds: (indexname * term option) list}; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
51 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
52 |
type cases = (string * T option) list; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
53 |
type cases_tactic = thm -> (cases * thm) Seq.seq; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
54 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
55 |
fun CASES cases tac st = Seq.map (pair cases) (tac st); |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
56 |
fun NO_CASES tac = CASES [] tac; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
57 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
58 |
fun SUBGOAL_CASES tac i st = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
59 |
(case try Logic.nth_prem (i, Thm.prop_of st) of |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
60 |
SOME goal => tac (goal, i) st |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
61 |
| NONE => Seq.empty); |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
62 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
63 |
fun (tac1 THEN_ALL_NEW_CASES tac2) i st = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
64 |
st |> tac1 i |> Seq.maps (fn (cases, st') => |
18237 | 65 |
CASES cases (Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st)) st'); |
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
66 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
67 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
68 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
69 |
(** consume facts **) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
70 |
|
18477 | 71 |
local |
72 |
||
73 |
fun unfold_prems n defs th = |
|
74 |
if null defs then th |
|
75 |
else Drule.fconv_rule (Drule.goals_conv (fn i => i <= n) (Tactic.rewrite true defs)) th; |
|
76 |
||
77 |
fun unfold_prems_concls defs th = |
|
78 |
if null defs orelse not (can Logic.dest_conjunction (Thm.concl_of th)) then th |
|
79 |
else |
|
80 |
Drule.fconv_rule |
|
81 |
(Drule.concl_conv ~1 (Drule.conjunction_conv ~1 |
|
82 |
(K (Drule.prems_conv ~1 (K (Tactic.rewrite true defs)))))) th; |
|
83 |
||
84 |
in |
|
85 |
||
86 |
fun consume defs facts ((xx, n), th) = |
|
18237 | 87 |
let val m = Int.min (length facts, n) in |
18477 | 88 |
th |
89 |
|> unfold_prems n defs |
|
90 |
|> unfold_prems_concls defs |
|
18237 | 91 |
|> Drule.multi_resolve (Library.take (m, facts)) |
18477 | 92 |
|> Seq.map (pair (xx, (n - m, Library.drop (m, facts)))) |
18237 | 93 |
end; |
10811 | 94 |
|
18477 | 95 |
end; |
96 |
||
10811 | 97 |
val consumes_tagN = "consumes"; |
98 |
||
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
99 |
fun lookup_consumes th = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
100 |
let fun err () = raise THM ("Malformed 'consumes' tag of theorem", 0, [th]) in |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
101 |
(case AList.lookup (op =) (Thm.tags_of_thm th) (consumes_tagN) of |
15531 | 102 |
NONE => NONE |
103 |
| SOME [s] => (case Syntax.read_nat s of SOME n => SOME n | _ => err ()) |
|
10530 | 104 |
| _ => err ()) |
105 |
end; |
|
106 |
||
18477 | 107 |
fun get_consumes th = the_default 0 (lookup_consumes th); |
108 |
||
15531 | 109 |
fun put_consumes NONE th = th |
110 |
| put_consumes (SOME n) th = th |
|
12761 | 111 |
|> Drule.untag_rule consumes_tagN |
112 |
|> Drule.tag_rule (consumes_tagN, [Library.string_of_int n]); |
|
10530 | 113 |
|
18477 | 114 |
fun add_consumes k th = put_consumes (SOME (k + get_consumes th)) th; |
115 |
||
12761 | 116 |
val save_consumes = put_consumes o lookup_consumes; |
117 |
||
15531 | 118 |
fun consumes n x = Drule.rule_attribute (K (put_consumes (SOME n))) x; |
16148
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
119 |
fun consumes_default n x = |
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
120 |
if Library.is_some (lookup_consumes (#2 x)) then x else consumes n x; |
8364 | 121 |
|
122 |
||
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
123 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
124 |
(** case names **) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
125 |
|
18237 | 126 |
val case_names_tagN = "case_names"; |
8364 | 127 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
128 |
fun add_case_names NONE = I |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
129 |
| add_case_names (SOME names) = |
18237 | 130 |
Drule.untag_rule case_names_tagN |
131 |
#> Drule.tag_rule (case_names_tagN, names); |
|
8364 | 132 |
|
18237 | 133 |
fun lookup_case_names th = AList.lookup (op =) (Thm.tags_of_thm th) case_names_tagN; |
10530 | 134 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
135 |
val save_case_names = add_case_names o lookup_case_names; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
136 |
val name = add_case_names o SOME; |
8427 | 137 |
fun case_names ss = Drule.rule_attribute (K (name ss)); |
8364 | 138 |
|
10530 | 139 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
140 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
141 |
(** case conclusions **) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
142 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
143 |
val case_concl_tagN = "case_conclusion"; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
144 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
145 |
fun is_case_concl name ((a, b :: _): tag) = (a = case_concl_tagN andalso b = name) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
146 |
| is_case_concl _ _ = false; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
147 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
148 |
fun add_case_concl (name, cs) = Drule.map_tags (fn tags => |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
149 |
filter_out (is_case_concl name) tags @ [(case_concl_tagN, name :: cs)]); |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
150 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
151 |
fun get_case_concls th name = |
18237 | 152 |
(case find_first (is_case_concl name) (Thm.tags_of_thm th) of |
153 |
SOME (_, _ :: cs) => cs |
|
154 |
| _ => []); |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
155 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
156 |
fun save_case_concls th = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
157 |
let val concls = Thm.tags_of_thm th |> List.mapPartial |
18237 | 158 |
(fn (a, b :: cs) => |
159 |
if a = case_concl_tagN then SOME (b, cs) else NONE |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
160 |
| _ => NONE) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
161 |
in fold add_case_concl concls end; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
162 |
|
18237 | 163 |
fun case_conclusion concl = Drule.rule_attribute (fn _ => add_case_concl concl); |
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
164 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
165 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
166 |
|
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
167 |
(** case declarations **) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
168 |
|
10530 | 169 |
(* access hints *) |
170 |
||
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
171 |
fun save th = save_consumes th #> save_case_names th #> save_case_concls th; |
12761 | 172 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
173 |
fun get th = |
12761 | 174 |
let |
18477 | 175 |
val n = get_consumes th; |
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
176 |
val cases = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
177 |
(case lookup_case_names th of |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
178 |
NONE => map (rpair [] o Library.string_of_int) (1 upto (Thm.nprems_of th - n)) |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
179 |
| SOME names => map (fn name => (name, get_case_concls th name)) names); |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
180 |
in (cases, n) end; |
10530 | 181 |
|
8364 | 182 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
183 |
(* extract cases *) |
8364 | 184 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
185 |
val case_conclN = "case"; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
186 |
val case_hypsN = "hyps"; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
187 |
val case_premsN = "prems"; |
13425
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
12809
diff
changeset
|
188 |
|
18375 | 189 |
val strip_params = map (apfst (perhaps (try Syntax.dest_skolem))) o Logic.strip_params; |
17167 | 190 |
|
18477 | 191 |
local |
192 |
||
18237 | 193 |
fun dest_binops cs tm = |
194 |
let |
|
195 |
val n = length cs; |
|
196 |
fun dest 0 _ = [] |
|
197 |
| dest 1 t = [t] |
|
198 |
| dest k (_ $ t $ u) = t :: dest (k - 1) u |
|
199 |
| dest _ _ = raise TERM ("Expected " ^ string_of_int n ^ " binop arguments", [tm]); |
|
200 |
in cs ~~ dest n tm end; |
|
201 |
||
18477 | 202 |
fun extract_cases is_open thy (split, raw_prop) name concls = |
8364 | 203 |
let |
18477 | 204 |
fun extract prop idx = |
205 |
let |
|
206 |
val xs = strip_params prop; |
|
207 |
val rename = if is_open then I else map (apfst Syntax.internal); |
|
208 |
val fixes = |
|
209 |
(case split of |
|
210 |
NONE => rename xs |
|
211 |
| SOME t => |
|
212 |
let val (us, vs) = splitAt (length (Logic.strip_params t), xs) |
|
213 |
in rename us @ vs end); |
|
214 |
fun abs_fixes t = Term.list_abs (fixes, t); |
|
215 |
val dest_conjuncts = map abs_fixes o List.concat o map Logic.dest_conjunctions; |
|
16148
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
216 |
|
18477 | 217 |
val asms = Logic.strip_assums_hyp prop; |
218 |
val assumes = |
|
219 |
(case split of |
|
220 |
NONE => [("", dest_conjuncts asms)] |
|
221 |
| SOME t => |
|
222 |
let val (hyps, prems) = splitAt (length (Logic.strip_assums_hyp t), asms) in |
|
223 |
[(case_hypsN, dest_conjuncts hyps), |
|
224 |
(case_premsN, dest_conjuncts prems)] |
|
225 |
end); |
|
16148
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
226 |
|
18477 | 227 |
val concl = ObjectLogic.drop_judgment thy (Logic.strip_assums_concl prop); |
228 |
val binds = (case_conclN, concl) :: dest_binops concls concl |
|
229 |
|> map (fn (x, t) => ((x, 0), SOME (abs_fixes t))); |
|
230 |
in (name ^ idx, SOME {fixes = fixes, assumes = assumes, binds = binds}) end; |
|
231 |
in |
|
232 |
(case Logic.dest_conjunctions (Drule.norm_hhf thy raw_prop) of |
|
233 |
[prop] => [extract prop ""] |
|
234 |
| props => map2 extract props (map string_of_int (1 upto length props))) |
|
235 |
end; |
|
16148
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
236 |
|
18477 | 237 |
in |
8364 | 238 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
239 |
fun make is_open split (thy, prop) cases = |
16148
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
240 |
let |
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
241 |
val n = length cases; |
16148
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
242 |
val nprems = Logic.count_prems (prop, 0); |
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
243 |
fun add_case (name, concls) (cs, i) = |
16148
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
244 |
((case try (fn () => |
5f15a14122dc
make: T option -- actually remove undefined cases;
wenzelm
parents:
15973
diff
changeset
|
245 |
(Option.map (curry Logic.nth_prem i) split, Logic.nth_prem (i, prop))) () of |
18477 | 246 |
NONE => [(name, NONE)] |
247 |
| SOME sp => extract_cases is_open thy sp name concls) @ cs, i - 1); |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
248 |
in fold_rev add_case (Library.drop (n - nprems, cases)) ([], n) |> #1 end; |
8364 | 249 |
|
18477 | 250 |
fun simple (thy, prop) name = |
251 |
extract_cases true thy (NONE, prop) name []; |
|
252 |
||
253 |
end; |
|
17361 | 254 |
|
8364 | 255 |
|
8427 | 256 |
(* params *) |
8364 | 257 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
258 |
fun rename_params xss th = |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
259 |
th |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
260 |
|> fold_index (fn (i, xs) => Thm.rename_params_rule (xs, i + 1)) xss |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
261 |
|> save th; |
8364 | 262 |
|
263 |
fun params xss = Drule.rule_attribute (K (rename_params xss)); |
|
264 |
||
18477 | 265 |
|
266 |
||
267 |
(** mutual_rule **) |
|
268 |
||
269 |
local |
|
270 |
||
271 |
fun equal_cterms ts us = |
|
272 |
list_ord (Term.fast_term_ord o pairself Thm.term_of) (ts, us) = EQUAL; |
|
273 |
||
274 |
fun prep_rule th = |
|
275 |
let |
|
276 |
val n = get_consumes th; |
|
277 |
val th' = Drule.freeze_all (Thm.permute_prems 0 n th); |
|
278 |
val prems = Library.take (Thm.nprems_of th' - n, Drule.cprems_of th'); |
|
279 |
val th'' = Drule.implies_elim_list th' (map Thm.assume prems); |
|
280 |
in (prems, (n, th'')) end; |
|
281 |
||
282 |
in |
|
283 |
||
284 |
fun mutual_rule [] = NONE |
|
285 |
| mutual_rule [th] = SOME ([0], th) |
|
286 |
| mutual_rule raw_rules = |
|
287 |
let |
|
288 |
val rules as (prems, _) :: _ = map prep_rule raw_rules; |
|
289 |
val (ns, ths) = split_list (map #2 rules); |
|
290 |
in |
|
291 |
if not (forall (equal_cterms prems o #1) rules) then NONE |
|
292 |
else |
|
293 |
SOME (ns, |
|
294 |
ths |
|
295 |
|> foldr1 (uncurry Drule.conj_intr) |
|
296 |
|> Drule.implies_intr_list prems |
|
297 |
|> Drule.standard' |
|
298 |
|> save (hd raw_rules) |
|
299 |
|> put_consumes (SOME 0)) |
|
300 |
end; |
|
301 |
||
302 |
end; |
|
303 |
||
8364 | 304 |
end; |
17113 | 305 |
|
18229
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
306 |
structure BasicRuleCases: BASIC_RULE_CASES = RuleCases; |
e5a624483a23
added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents:
18188
diff
changeset
|
307 |
open BasicRuleCases; |