6442
|
1 |
(* Title: HOL/Tools/induct_method.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Markus Wenzel, TU Muenchen
|
|
4 |
|
8308
|
5 |
Proof by cases and induction on types (intro) and sets (elim).
|
6442
|
6 |
*)
|
|
7 |
|
|
8 |
signature INDUCT_METHOD =
|
|
9 |
sig
|
8308
|
10 |
val print_global_rules: theory -> unit
|
|
11 |
val print_local_rules: Proof.context -> unit
|
|
12 |
val cases_type_global: string -> theory attribute
|
|
13 |
val cases_set_global: string -> theory attribute
|
|
14 |
val cases_type_local: string -> Proof.context attribute
|
|
15 |
val cases_set_local: string -> Proof.context attribute
|
|
16 |
val induct_type_global: string -> theory attribute
|
|
17 |
val induct_set_global: string -> theory attribute
|
|
18 |
val induct_type_local: string -> Proof.context attribute
|
|
19 |
val induct_set_local: string -> Proof.context attribute
|
6442
|
20 |
val setup: (theory -> theory) list
|
|
21 |
end;
|
|
22 |
|
|
23 |
structure InductMethod: INDUCT_METHOD =
|
|
24 |
struct
|
|
25 |
|
8308
|
26 |
(** global and local induct data **)
|
6442
|
27 |
|
8308
|
28 |
(* rules *)
|
|
29 |
|
|
30 |
type rules = (string * thm) NetRules.T;
|
|
31 |
|
|
32 |
fun eq_rule ((s1:string, th1), (s2, th2)) = s1 = s2 andalso Thm.eq_thm (th1, th2);
|
|
33 |
|
|
34 |
val type_rules = NetRules.init eq_rule (Thm.concl_of o #2);
|
|
35 |
val set_rules = NetRules.init eq_rule (Thm.major_prem_of o #2);
|
|
36 |
|
|
37 |
fun lookup_rule (rs:rules) name = Library.assoc (NetRules.rules rs, name);
|
|
38 |
|
|
39 |
fun print_rules kind rs =
|
|
40 |
let val thms = map snd (NetRules.rules rs)
|
|
41 |
in Pretty.writeln (Pretty.big_list (kind ^ " rules:") (map Display.pretty_thm thms)) end;
|
|
42 |
|
|
43 |
|
|
44 |
(* theory data kind 'HOL/induct_method' *)
|
|
45 |
|
|
46 |
structure GlobalInductArgs =
|
|
47 |
struct
|
|
48 |
val name = "HOL/induct_method";
|
|
49 |
type T = (rules * rules) * (rules * rules);
|
|
50 |
|
|
51 |
val empty = ((type_rules, set_rules), (type_rules, set_rules));
|
|
52 |
val copy = I;
|
|
53 |
val prep_ext = I;
|
|
54 |
fun merge (((casesT1, casesS1), (inductT1, inductS1)),
|
|
55 |
((casesT2, casesS2), (inductT2, inductS2))) =
|
|
56 |
((NetRules.merge (casesT1, casesT2), NetRules.merge (casesS1, casesS2)),
|
|
57 |
(NetRules.merge (inductT1, inductT2), NetRules.merge (inductS1, inductS2)));
|
|
58 |
|
|
59 |
fun print _ ((casesT, casesS), (inductT, inductS)) =
|
|
60 |
(print_rules "type cases" casesT;
|
|
61 |
print_rules "set cases" casesS;
|
|
62 |
print_rules "type induct" inductT;
|
|
63 |
print_rules "set induct" inductS);
|
|
64 |
end;
|
|
65 |
|
|
66 |
structure GlobalInduct = TheoryDataFun(GlobalInductArgs);
|
|
67 |
val print_global_rules = GlobalInduct.print;
|
|
68 |
|
|
69 |
|
|
70 |
(* proof data kind 'HOL/induct_method' *)
|
|
71 |
|
|
72 |
structure LocalInductArgs =
|
|
73 |
struct
|
|
74 |
val name = "HOL/induct_method";
|
|
75 |
type T = GlobalInductArgs.T;
|
8278
|
76 |
|
8308
|
77 |
fun init thy = GlobalInduct.get thy;
|
|
78 |
fun print x = GlobalInductArgs.print x;
|
|
79 |
end;
|
|
80 |
|
|
81 |
structure LocalInduct = ProofDataFun(LocalInductArgs);
|
|
82 |
val print_local_rules = LocalInduct.print;
|
|
83 |
|
|
84 |
|
|
85 |
(* access rules *)
|
|
86 |
|
|
87 |
val get_cases = #1 o LocalInduct.get;
|
|
88 |
val get_induct = #2 o LocalInduct.get;
|
|
89 |
|
|
90 |
val lookup_casesT = lookup_rule o #1 o get_cases;
|
|
91 |
val lookup_casesS = lookup_rule o #2 o get_cases;
|
|
92 |
val lookup_inductT = lookup_rule o #1 o get_induct;
|
|
93 |
val lookup_inductS = lookup_rule o #2 o get_induct;
|
|
94 |
|
|
95 |
|
|
96 |
|
|
97 |
(** attributes **)
|
|
98 |
|
|
99 |
local
|
|
100 |
|
|
101 |
fun mk_att f g name (x, thm) = (f (g (name, thm)) x, thm);
|
|
102 |
|
|
103 |
fun add_casesT rule x = apfst (apfst (NetRules.insert rule)) x;
|
|
104 |
fun add_casesS rule x = apfst (apsnd (NetRules.insert rule)) x;
|
|
105 |
fun add_inductT rule x = apsnd (apfst (NetRules.insert rule)) x;
|
|
106 |
fun add_inductS rule x = apsnd (apsnd (NetRules.insert rule)) x;
|
|
107 |
|
|
108 |
in
|
|
109 |
|
|
110 |
val cases_type_global = mk_att GlobalInduct.map add_casesT;
|
|
111 |
val cases_set_global = mk_att GlobalInduct.map add_casesS;
|
|
112 |
val induct_type_global = mk_att GlobalInduct.map add_inductT;
|
|
113 |
val induct_set_global = mk_att GlobalInduct.map add_inductS;
|
|
114 |
|
|
115 |
val cases_type_local = mk_att LocalInduct.map add_casesT;
|
|
116 |
val cases_set_local = mk_att LocalInduct.map add_casesS;
|
|
117 |
val induct_type_local = mk_att LocalInduct.map add_inductT;
|
|
118 |
val induct_set_local = mk_att LocalInduct.map add_inductS;
|
|
119 |
|
|
120 |
end;
|
|
121 |
|
|
122 |
|
|
123 |
|
|
124 |
(** misc utils **)
|
8278
|
125 |
|
|
126 |
fun vars_of tm = (*ordered left-to-right, preferring right!*)
|
8308
|
127 |
Term.foldl_aterms (fn (ts, t as Var _) => t :: ts | (ts, _) => ts) ([], tm)
|
8278
|
128 |
|> Library.distinct |> rev;
|
|
129 |
|
8308
|
130 |
fun type_name t =
|
|
131 |
#1 (Term.dest_Type (Term.type_of t))
|
|
132 |
handle TYPE _ => raise TERM ("Bad type of term argument", [t]);
|
8278
|
133 |
|
|
134 |
|
|
135 |
|
|
136 |
(** cases method **)
|
|
137 |
|
8308
|
138 |
(*
|
|
139 |
rule selection:
|
|
140 |
cases - classical case split
|
|
141 |
<x:A> cases - set elimination
|
|
142 |
... cases t - datatype exhaustion
|
|
143 |
... cases ... r - explicit rule
|
|
144 |
*)
|
8278
|
145 |
|
8308
|
146 |
fun cases_var thm =
|
|
147 |
(case try (hd o vars_of o hd o Logic.strip_assums_hyp o Library.last_elem o Thm.prems_of) thm of
|
|
148 |
None => raise THM ("Malformed cases rule", 0, [thm])
|
|
149 |
| Some x => x);
|
|
150 |
|
|
151 |
fun cases_tac (ctxt, args) facts =
|
|
152 |
let
|
|
153 |
val sg = ProofContext.sign_of ctxt;
|
|
154 |
val cert = Thm.cterm_of sg;
|
8278
|
155 |
|
8308
|
156 |
fun inst_rule t thm =
|
|
157 |
Drule.cterm_instantiate [(cert (cases_var thm), cert t)] thm;
|
6442
|
158 |
|
8308
|
159 |
val thms =
|
|
160 |
(case (args, facts) of
|
|
161 |
((None, None), []) => [case_split_thm]
|
|
162 |
| ((None, None), th :: _) =>
|
|
163 |
NetRules.may_unify (#2 (get_cases ctxt))
|
|
164 |
(Logic.strip_assums_concl (#prop (Thm.rep_thm th)))
|
|
165 |
|> map #2
|
|
166 |
| ((Some t, None), _) =>
|
|
167 |
let val name = type_name t in
|
|
168 |
(case lookup_casesT ctxt name of
|
|
169 |
None => error ("No cases rule for type: " ^ quote name)
|
|
170 |
| Some thm => [inst_rule t thm])
|
|
171 |
end
|
|
172 |
| ((None, Some thm), _) => [thm]
|
|
173 |
| ((Some t, Some thm), _) => [inst_rule t thm]);
|
|
174 |
in Method.rule_tac thms facts end;
|
8278
|
175 |
|
8308
|
176 |
val cases_meth = Method.METHOD o (FINDGOAL oo cases_tac);
|
8278
|
177 |
|
|
178 |
|
|
179 |
|
|
180 |
(** induct method **)
|
|
181 |
|
8308
|
182 |
(*
|
|
183 |
rule selection:
|
|
184 |
induct - mathematical induction
|
|
185 |
<x:A> induct - set induction
|
|
186 |
... induct x - datatype induction
|
|
187 |
... induct ... r - explicit rule
|
|
188 |
*)
|
8278
|
189 |
|
8308
|
190 |
fun induct_tac (ctxt, args) facts =
|
|
191 |
let
|
|
192 |
val sg = ProofContext.sign_of ctxt;
|
|
193 |
val cert = Thm.cterm_of sg;
|
|
194 |
|
|
195 |
fun prep_inst (concl, ts) =
|
|
196 |
let val xs = vars_of concl; val n = length xs - length ts in
|
|
197 |
if n < 0 then error "More arguments given than in induction rule"
|
|
198 |
else map cert (Library.drop (n, xs)) ~~ map cert ts
|
|
199 |
end;
|
8278
|
200 |
|
8308
|
201 |
fun inst_rule insts thm =
|
|
202 |
Drule.cterm_instantiate (flat (map2 prep_inst
|
|
203 |
(HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of thm)), insts))) thm;
|
8278
|
204 |
|
8308
|
205 |
val thms =
|
|
206 |
(case (args, facts) of
|
|
207 |
(([], None), []) => [nat_induct]
|
|
208 |
| (([], None), th :: _) =>
|
|
209 |
NetRules.may_unify (#2 (get_induct ctxt))
|
|
210 |
(Logic.strip_assums_concl (#prop (Thm.rep_thm th)))
|
|
211 |
|> map #2
|
|
212 |
| ((insts, None), _) =>
|
|
213 |
let val name = type_name (last_elem (hd insts)) in
|
|
214 |
(case lookup_inductT ctxt name of
|
|
215 |
None => error ("No induct rule for type: " ^ quote name)
|
|
216 |
| Some thm => [inst_rule insts thm])
|
|
217 |
end
|
|
218 |
| (([], Some thm), _) => [thm]
|
|
219 |
| ((insts, Some thm), _) => [inst_rule insts thm]);
|
|
220 |
in Method.rule_tac thms facts end;
|
8278
|
221 |
|
8308
|
222 |
val induct_meth = Method.METHOD o (FINDGOAL oo induct_tac);
|
8278
|
223 |
|
|
224 |
|
|
225 |
|
|
226 |
(** concrete syntax **)
|
|
227 |
|
8308
|
228 |
val casesN = "cases";
|
|
229 |
val inductN = "induct";
|
|
230 |
val typeN = "type";
|
|
231 |
val setN = "set";
|
|
232 |
val ruleN = "rule";
|
|
233 |
|
|
234 |
|
|
235 |
(* attributes *)
|
|
236 |
|
|
237 |
fun spec k = (Args.$$$ k -- Args.$$$ ":") |-- Args.!!! Args.name;
|
|
238 |
|
|
239 |
fun attrib sign_of add_type add_set = Scan.depend (fn x =>
|
|
240 |
let val sg = sign_of x in
|
|
241 |
spec typeN >> (add_type o Sign.intern_tycon sg) ||
|
|
242 |
spec setN >> (add_set o Sign.intern_const sg)
|
|
243 |
end >> pair x);
|
|
244 |
|
|
245 |
val cases_attr =
|
|
246 |
(Attrib.syntax (attrib Theory.sign_of cases_type_global cases_set_global),
|
|
247 |
Attrib.syntax (attrib ProofContext.sign_of cases_type_local cases_set_local));
|
|
248 |
|
|
249 |
val induct_attr =
|
|
250 |
(Attrib.syntax (attrib Theory.sign_of induct_type_global induct_set_global),
|
|
251 |
Attrib.syntax (attrib ProofContext.sign_of induct_type_local induct_set_local));
|
|
252 |
|
|
253 |
|
|
254 |
(* methods *)
|
|
255 |
|
8278
|
256 |
local
|
6442
|
257 |
|
8308
|
258 |
fun err k get name =
|
|
259 |
(case get name of Some x => x
|
|
260 |
| None => error ("No rule for " ^ k ^ " " ^ quote name));
|
6442
|
261 |
|
8308
|
262 |
fun rule get_type get_set =
|
|
263 |
Scan.depend (fn ctxt =>
|
|
264 |
let val sg = ProofContext.sign_of ctxt in
|
|
265 |
spec typeN >> (err typeN (get_type ctxt) o Sign.intern_tycon sg) ||
|
|
266 |
spec setN >> (err setN (get_set ctxt) o Sign.intern_const sg)
|
|
267 |
end >> pair ctxt) ||
|
|
268 |
Scan.lift (Args.$$$ ruleN -- Args.$$$ ":") |-- Attrib.local_thm;
|
6442
|
269 |
|
8308
|
270 |
val cases_rule = rule lookup_casesT lookup_casesS;
|
|
271 |
val induct_rule = rule lookup_inductT lookup_inductS;
|
6442
|
272 |
|
8308
|
273 |
val kind = (Args.$$$ typeN || Args.$$$ setN || Args.$$$ ruleN) -- Args.$$$ ":";
|
|
274 |
val term = Scan.unless (Scan.lift kind) Args.local_term;
|
6446
|
275 |
|
8278
|
276 |
in
|
|
277 |
|
8308
|
278 |
val cases_args = Method.syntax (Scan.option term -- Scan.option cases_rule);
|
|
279 |
val induct_args = Method.syntax (Args.and_list (Scan.repeat1 term) -- Scan.option induct_rule);
|
8278
|
280 |
|
|
281 |
end;
|
6446
|
282 |
|
|
283 |
|
6442
|
284 |
|
8278
|
285 |
(** theory setup **)
|
6446
|
286 |
|
8278
|
287 |
val setup =
|
8308
|
288 |
[GlobalInduct.init, LocalInduct.init,
|
|
289 |
Attrib.add_attributes
|
|
290 |
[(casesN, cases_attr, "cases rule for type or set"),
|
|
291 |
(inductN, induct_attr, "induction rule for type or set")],
|
|
292 |
Method.add_methods
|
|
293 |
[("cases", cases_meth oo cases_args, "case analysis on types or sets"),
|
|
294 |
("induct", induct_meth oo induct_args, "induction on types or sets")]];
|
6442
|
295 |
|
|
296 |
end;
|