src/Tools/induction.ML
author wenzelm
Sun, 13 Dec 2015 21:56:15 +0100
changeset 61841 4d3527b94f2a
parent 59970 e9f73d87d904
child 61844 007d3b34080f
permissions -rw-r--r--
more general types Proof.method / context_tactic; proper context for Method.insert_tac; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59931
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
     1
(*  Title:      Tools/induction.ML
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
     2
    Author:     Tobias Nipkow, TU Muenchen
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
     3
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
     4
Alternative proof method "induction" that gives induction hypotheses the
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
     5
name "IH".
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
     6
*)
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
     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
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 59970
diff changeset
    10
  val induction_tac: bool -> (binding option * (term * bool)) option list list ->
45014
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    11
    (string * typ) list list -> term option list -> thm list option ->
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 59970
diff changeset
    12
    thm list -> int -> context_tactic
45014
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    13
end
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    14
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    15
structure Induction: INDUCTION =
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    16
struct
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    17
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    18
val ind_hypsN = "IH";
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    19
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    20
fun preds_of t =
59931
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    21
  (case strip_comb t of
45014
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    22
    (p as Var _, _) => [p]
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    23
  | (p as Free _, _) => [p]
59931
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    24
  | (_, ts) => maps preds_of ts);
45014
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    25
56506
wenzelm
parents: 45014
diff changeset
    26
fun name_hyps (arg as ((cases, consumes), th)) =
59931
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    27
  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
    28
  else
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    29
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58826
diff changeset
    30
      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
    31
      val prems' = drop consumes prems;
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    32
      val ps = preds_of concl;
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    33
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    34
      fun hname_of t =
59931
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    35
        if exists_subterm (member (op aconv) ps) t
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    36
        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
    37
59931
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    38
      val hnamess = map (map hname_of o Logic.strip_assums_hyp) prems';
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    39
      val n = Int.min (length hnamess, length cases);
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    40
      val cases' =
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    41
        map (fn (((cn, _), concls), hns) => ((cn, hns), concls))
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    42
          (take n cases ~~ take n hnamess);
5ec4f97dd6d4 proper header;
wenzelm
parents: 59582
diff changeset
    43
    in ((cases', consumes), th) end;
45014
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    44
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59940
diff changeset
    45
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
    46
59940
087d81f5213e local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents: 59931
diff changeset
    47
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
    48
0e847655b2d8 New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
diff changeset
    49
end