author | wenzelm |
Mon, 08 Jun 2015 22:04:19 +0200 | |
changeset 60393 | b640770117fd |
parent 59970 | e9f73d87d904 |
child 61841 | 4d3527b94f2a |
permissions | -rw-r--r-- |
59931 | 1 |
(* Title: Tools/induction.ML |
2 |
Author: Tobias Nipkow, TU Muenchen |
|
3 |
||
4 |
Alternative proof method "induction" that gives induction hypotheses the |
|
5 |
name "IH". |
|
6 |
*) |
|
7 |
||
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
8 |
signature INDUCTION = |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
9 |
sig |
59931 | 10 |
val induction_tac: Proof.context -> bool -> |
11 |
(binding option * (term * bool)) option list list -> |
|
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
12 |
(string * typ) list list -> term option list -> thm list option -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
13 |
thm list -> int -> cases_tactic |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
14 |
end |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
15 |
|
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
16 |
structure Induction: INDUCTION = |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
17 |
struct |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
18 |
|
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
19 |
val ind_hypsN = "IH"; |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
20 |
|
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
21 |
fun preds_of t = |
59931 | 22 |
(case strip_comb t of |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
23 |
(p as Var _, _) => [p] |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
24 |
| (p as Free _, _) => [p] |
59931 | 25 |
| (_, ts) => maps preds_of ts); |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
26 |
|
56506 | 27 |
fun name_hyps (arg as ((cases, consumes), th)) = |
59931 | 28 |
if not (forall (null o #2 o #1) cases) then arg |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
29 |
else |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
30 |
let |
59582 | 31 |
val (prems, concl) = Logic.strip_horn (Thm.prop_of th); |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
32 |
val prems' = drop consumes prems; |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
33 |
val ps = preds_of concl; |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
34 |
|
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
35 |
fun hname_of t = |
59931 | 36 |
if exists_subterm (member (op aconv) ps) t |
37 |
then ind_hypsN else Rule_Cases.case_hypsN; |
|
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
38 |
|
59931 | 39 |
val hnamess = map (map hname_of o Logic.strip_assums_hyp) prems'; |
40 |
val n = Int.min (length hnamess, length cases); |
|
41 |
val cases' = |
|
42 |
map (fn (((cn, _), concls), hns) => ((cn, hns), concls)) |
|
43 |
(take n cases ~~ take n hnamess); |
|
44 |
in ((cases', consumes), th) end; |
|
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
45 |
|
59970 | 46 |
val induction_tac = Induct.gen_induct_tac name_hyps; |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
47 |
|
59940
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents:
59931
diff
changeset
|
48 |
val _ = Theory.local_setup (Induct.gen_induct_setup @{binding induction} induction_tac); |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
49 |
|
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff
changeset
|
50 |
end |