src/HOL/Tools/induct_method.ML
author wenzelm
Tue, 07 Sep 1999 18:10:33 +0200
changeset 7512 930e5947562d
parent 7017 e4e64a0b0b6b
child 8154 dab09e1ad594
permissions -rw-r--r--
rule option;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Tools/induct_method.ML
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     4
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     5
Proof by induction on types / set / functions.
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     6
*)
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     7
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     8
signature INDUCT_METHOD =
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
     9
sig
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    10
  val setup: (theory -> theory) list
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    11
end;
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    12
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    13
structure InductMethod: INDUCT_METHOD =
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    14
struct
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    15
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    16
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    17
(* type induct_kind *)
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    18
7512
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    19
datatype induct_kind = Type | Set | Function | Rule;
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    20
7512
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    21
val kind_name =
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    22
  Args.$$$ "type" >> K Type ||
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    23
  Args.$$$ "set" >> K Set ||
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    24
  Args.$$$ "function" >> K Function ||
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    25
  Args.$$$ "rule" >> K Rule;
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    26
7512
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    27
val kind = kind_name --| Args.$$$ ":";
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    28
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    29
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    30
(* induct_rule *)
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    31
7017
e4e64a0b0b6b Replaced datatype_info by datatype_info_err.
berghofe
parents: 6518
diff changeset
    32
fun kind_rule Type = (#induction oo DatatypePackage.datatype_info_err, Sign.intern_tycon)
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    33
  | kind_rule Set = ((#induct o #2) oo InductivePackage.get_inductive, Sign.intern_const)
7512
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    34
  | kind_rule Function = (#induct oo RecdefPackage.get_recdef, Sign.intern_const)
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    35
  | kind_rule Rule = (PureThy.get_thm, K I);
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    36
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    37
fun induct_rule thy kind name =
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    38
  let val (ind_rule, intern) = kind_rule kind
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    39
  in ind_rule thy (intern (Theory.sign_of thy) name) end;
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    40
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    41
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    42
(* induct_tac *)
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    43
6518
e9d6f165f9c1 support forward chaining;
wenzelm
parents: 6446
diff changeset
    44
fun induct_tac thy insts opt_kind_name facts i state =
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    45
  let
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    46
    val (kind, name) =
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    47
      (case opt_kind_name of Some kind_name => kind_name | None =>
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    48
        (case try (#1 o Term.dest_Type o Term.type_of o #2 o hd) insts of
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    49
          Some name => (Type, name)
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    50
        | None => error "Unable to figure out induction rule"));
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    51
    val rule = Drule.incr_indexes_wrt [] [] [] [state] (induct_rule thy kind name);
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    52
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    53
    val cert = Thm.cterm_of (Sign.merge (pairself Thm.sign_of_thm (state, rule)));
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    54
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    55
    fun prep_inst ((P $ x), (Some Q, y)) = [(cert P, cert Q), (cert x, cert y)]
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    56
      | prep_inst ((_ $ x), (None, y)) = [(cert x, cert y)]
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    57
      | prep_inst _ = raise THM ("Malformed induction rule", 0, [rule]);
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    58
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    59
    val prep_insts = flat o map2 prep_inst;
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    60
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    61
    val inst_rule =
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    62
      if null insts then rule
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    63
      else Drule.cterm_instantiate (prep_insts
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    64
        (DatatypeAux.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of rule)), insts)) rule;
6518
e9d6f165f9c1 support forward chaining;
wenzelm
parents: 6446
diff changeset
    65
  in Method.rule_tac [inst_rule] facts i state end;
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    66
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    67
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    68
(* induct_method *)
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    69
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    70
val term = Scan.unless (Scan.lift kind) Args.local_term;
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    71
val inst = term -- term >> (apfst Some) || term >> pair None;
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    72
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    73
fun induct ctxt =
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
    74
  Args.and_list inst -- Scan.option (Scan.lift (kind -- Args.name))
6518
e9d6f165f9c1 support forward chaining;
wenzelm
parents: 6446
diff changeset
    75
  >> (fn (is, k) => Method.METHOD (FIRSTGOAL o induct_tac (ProofContext.theory_of ctxt) is k));
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    76
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    77
fun induct_meth src ctxt = #2 (Method.syntax (induct ctxt) ctxt src);
7512
930e5947562d rule option;
wenzelm
parents: 7017
diff changeset
    78
val induct_method = ("induct", induct_meth, "induction on types, sets, functions etc.");
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    79
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    80
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    81
(* theory setup *)
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    82
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    83
val setup = [Method.add_methods [induct_method]];
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    84
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    85
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    86
end;