wenzelm@59931: (* Title: Tools/induction.ML wenzelm@59931: Author: Tobias Nipkow, TU Muenchen wenzelm@59931: wenzelm@59931: Alternative proof method "induction" that gives induction hypotheses the wenzelm@59931: name "IH". wenzelm@59931: *) wenzelm@59931: nipkow@45014: signature INDUCTION = nipkow@45014: sig wenzelm@59931: val induction_tac: Proof.context -> bool -> wenzelm@59931: (binding option * (term * bool)) option list list -> nipkow@45014: (string * typ) list list -> term option list -> thm list option -> nipkow@45014: thm list -> int -> cases_tactic nipkow@45014: end nipkow@45014: nipkow@45014: structure Induction: INDUCTION = nipkow@45014: struct nipkow@45014: nipkow@45014: val ind_hypsN = "IH"; nipkow@45014: nipkow@45014: fun preds_of t = wenzelm@59931: (case strip_comb t of nipkow@45014: (p as Var _, _) => [p] nipkow@45014: | (p as Free _, _) => [p] wenzelm@59931: | (_, ts) => maps preds_of ts); nipkow@45014: wenzelm@56506: fun name_hyps (arg as ((cases, consumes), th)) = wenzelm@59931: if not (forall (null o #2 o #1) cases) then arg nipkow@45014: else nipkow@45014: let wenzelm@59582: val (prems, concl) = Logic.strip_horn (Thm.prop_of th); nipkow@45014: val prems' = drop consumes prems; nipkow@45014: val ps = preds_of concl; nipkow@45014: nipkow@45014: fun hname_of t = wenzelm@59931: if exists_subterm (member (op aconv) ps) t wenzelm@59931: then ind_hypsN else Rule_Cases.case_hypsN; nipkow@45014: wenzelm@59931: val hnamess = map (map hname_of o Logic.strip_assums_hyp) prems'; wenzelm@59931: val n = Int.min (length hnamess, length cases); wenzelm@59931: val cases' = wenzelm@59931: map (fn (((cn, _), concls), hns) => ((cn, hns), concls)) wenzelm@59931: (take n cases ~~ take n hnamess); wenzelm@59931: in ((cases', consumes), th) end; nipkow@45014: wenzelm@59931: val induction_tac = Induct.gen_induct_tac (K name_hyps); nipkow@45014: wenzelm@59940: val _ = Theory.local_setup (Induct.gen_induct_setup @{binding induction} induction_tac); nipkow@45014: nipkow@45014: end