| author | wenzelm | 
| Fri, 06 Mar 2015 00:00:57 +0100 | |
| changeset 59617 | b60e65ad13df | 
| parent 59582 | 0fbed69ff081 | 
| child 59931 | 5ec4f97dd6d4 | 
| permissions | -rw-r--r-- | 
| 
45014
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
1  | 
signature INDUCTION =  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
2  | 
sig  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
3  | 
val induction_tac: Proof.context -> bool -> (binding option * (term * bool)) option list list ->  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
4  | 
(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
 | 
5  | 
thm list -> int -> cases_tactic  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
6  | 
end  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
8  | 
structure Induction: INDUCTION =  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
9  | 
struct  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
11  | 
val ind_hypsN = "IH";  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
13  | 
fun preds_of t =  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
14  | 
(case strip_comb t of  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
15  | 
(p as Var _, _) => [p]  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
16  | 
| (p as Free _, _) => [p]  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
17  | 
| (_, ts) => flat(map preds_of ts))  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
18  | 
|
| 56506 | 19  | 
fun name_hyps (arg as ((cases, consumes), th)) =  | 
| 
45014
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
20  | 
if not(forall (null o #2 o #1) cases) then arg  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
21  | 
else  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
22  | 
let  | 
| 59582 | 23  | 
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
 | 
24  | 
val prems' = drop consumes prems;  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
25  | 
val ps = preds_of concl;  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
27  | 
fun hname_of t =  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
28  | 
if exists_subterm (member (op =) ps) t  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
29  | 
then ind_hypsN else Rule_Cases.case_hypsN  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
31  | 
val hnamess = map (map hname_of o Logic.strip_assums_hyp) prems'  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
32  | 
val n = Int.min (length hnamess, length cases)  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
33  | 
val cases' = map (fn (((cn,_),concls),hns) => ((cn,hns),concls))  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
34  | 
(take n cases ~~ take n hnamess)  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
35  | 
in ((cases',consumes),th) end  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
36  | 
|
| 56506 | 37  | 
val induction_tac = Induct.gen_induct_tac (K name_hyps)  | 
| 
45014
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
38  | 
|
| 58826 | 39  | 
val _ = Theory.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
 | 
40  | 
|
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
41  | 
end  | 
| 
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents:  
diff
changeset
 | 
42  |