src/HOL/Tools/induct_method.ML
author wenzelm
Sun, 27 Feb 2000 15:33:35 +0100
changeset 8308 45e11d3ccbe4
parent 8295 ed9fc488b980
child 8315 c9b4f1c47816
permissions -rw-r--r--
cases/induct attributes; use NetRules for storage; tuned rule selection; tuned concrete syntax;
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
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
     5
Proof by cases and induction on types (intro) and sets (elim).
6442
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
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    10
  val print_global_rules: theory -> unit
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    11
  val print_local_rules: Proof.context -> unit
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    12
  val cases_type_global: string -> theory attribute
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    13
  val cases_set_global: string -> theory attribute
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    14
  val cases_type_local: string -> Proof.context attribute
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    15
  val cases_set_local: string -> Proof.context attribute
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    16
  val induct_type_global: string -> theory attribute
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    17
  val induct_set_global: string -> theory attribute
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    18
  val induct_type_local: string -> Proof.context attribute
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    19
  val induct_set_local: string -> Proof.context attribute
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    20
  val setup: (theory -> theory) list
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    21
end;
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    22
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    23
structure InductMethod: INDUCT_METHOD =
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    24
struct
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    25
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    26
(** global and local induct data **)
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
    27
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    28
(* rules *)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    29
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    30
type rules = (string * thm) NetRules.T;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    31
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    32
fun eq_rule ((s1:string, th1), (s2, th2)) = s1 = s2 andalso Thm.eq_thm (th1, th2);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    33
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    34
val type_rules = NetRules.init eq_rule (Thm.concl_of o #2);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    35
val set_rules = NetRules.init eq_rule (Thm.major_prem_of o #2);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    36
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    37
fun lookup_rule (rs:rules) name = Library.assoc (NetRules.rules rs, name);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    38
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    39
fun print_rules kind rs =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    40
  let val thms = map snd (NetRules.rules rs)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    41
  in Pretty.writeln (Pretty.big_list (kind ^ " rules:") (map Display.pretty_thm thms)) end;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    42
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    43
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    44
(* theory data kind 'HOL/induct_method' *)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    45
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    46
structure GlobalInductArgs =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    47
struct
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    48
  val name = "HOL/induct_method";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    49
  type T = (rules * rules) * (rules * rules);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    50
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    51
  val empty = ((type_rules, set_rules), (type_rules, set_rules));
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    52
  val copy = I;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    53
  val prep_ext = I;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    54
  fun merge (((casesT1, casesS1), (inductT1, inductS1)),
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    55
      ((casesT2, casesS2), (inductT2, inductS2))) =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    56
    ((NetRules.merge (casesT1, casesT2), NetRules.merge (casesS1, casesS2)),
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    57
      (NetRules.merge (inductT1, inductT2), NetRules.merge (inductS1, inductS2)));
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    58
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    59
  fun print _ ((casesT, casesS), (inductT, inductS)) =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    60
    (print_rules "type cases" casesT;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    61
      print_rules "set cases" casesS;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    62
      print_rules "type induct" inductT;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    63
      print_rules "set induct" inductS);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    64
end;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    65
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    66
structure GlobalInduct = TheoryDataFun(GlobalInductArgs);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    67
val print_global_rules = GlobalInduct.print;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    68
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    69
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    70
(* proof data kind 'HOL/induct_method' *)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    71
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    72
structure LocalInductArgs =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    73
struct
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    74
  val name = "HOL/induct_method";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    75
  type T = GlobalInductArgs.T;
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
    76
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    77
  fun init thy = GlobalInduct.get thy;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    78
  fun print x = GlobalInductArgs.print x;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    79
end;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    80
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    81
structure LocalInduct = ProofDataFun(LocalInductArgs);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    82
val print_local_rules = LocalInduct.print;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    83
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    84
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    85
(* access rules *)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    86
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    87
val get_cases = #1 o LocalInduct.get;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    88
val get_induct = #2 o LocalInduct.get;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    89
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    90
val lookup_casesT = lookup_rule o #1 o get_cases;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    91
val lookup_casesS = lookup_rule o #2 o get_cases;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    92
val lookup_inductT = lookup_rule o #1 o get_induct;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    93
val lookup_inductS = lookup_rule o #2 o get_induct;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    94
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    95
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    96
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    97
(** attributes **)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    98
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
    99
local
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   100
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   101
fun mk_att f g name (x, thm) = (f (g (name, thm)) x, thm);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   102
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   103
fun add_casesT rule x = apfst (apfst (NetRules.insert rule)) x;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   104
fun add_casesS rule x = apfst (apsnd (NetRules.insert rule)) x;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   105
fun add_inductT rule x = apsnd (apfst (NetRules.insert rule)) x;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   106
fun add_inductS rule x = apsnd (apsnd (NetRules.insert rule)) x;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   107
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   108
in
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   109
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   110
val cases_type_global = mk_att GlobalInduct.map add_casesT;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   111
val cases_set_global = mk_att GlobalInduct.map add_casesS;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   112
val induct_type_global = mk_att GlobalInduct.map add_inductT;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   113
val induct_set_global = mk_att GlobalInduct.map add_inductS;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   114
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   115
val cases_type_local = mk_att LocalInduct.map add_casesT;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   116
val cases_set_local = mk_att LocalInduct.map add_casesS;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   117
val induct_type_local = mk_att LocalInduct.map add_inductT;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   118
val induct_set_local = mk_att LocalInduct.map add_inductS;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   119
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   120
end;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   121
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   122
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   123
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   124
(** misc utils **)
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   125
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   126
fun vars_of tm =        (*ordered left-to-right, preferring right!*)
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   127
  Term.foldl_aterms (fn (ts, t as Var _) => t :: ts | (ts, _) => ts) ([], tm)
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   128
  |> Library.distinct |> rev;
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   129
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   130
fun type_name t =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   131
  #1 (Term.dest_Type (Term.type_of t))
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   132
    handle TYPE _ => raise TERM ("Bad type of term argument", [t]);
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   133
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   134
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   135
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   136
(** cases method **)
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   137
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   138
(*
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   139
  rule selection:
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   140
        cases         - classical case split
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   141
  <x:A> cases         - set elimination
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   142
  ...   cases t       - datatype exhaustion
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   143
  ...   cases ... r   - explicit rule
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   144
*)
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   145
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   146
fun cases_var thm =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   147
  (case try (hd o vars_of o hd o Logic.strip_assums_hyp o Library.last_elem o Thm.prems_of) thm of
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   148
    None => raise THM ("Malformed cases rule", 0, [thm])
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   149
  | Some x => x);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   150
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   151
fun cases_tac (ctxt, args) facts =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   152
  let
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   153
    val sg = ProofContext.sign_of ctxt;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   154
    val cert = Thm.cterm_of sg;
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   155
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   156
    fun inst_rule t thm =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   157
      Drule.cterm_instantiate [(cert (cases_var thm), cert t)] thm;
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   158
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   159
    val thms =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   160
      (case (args, facts) of
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   161
        ((None, None), []) => [case_split_thm]
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   162
      | ((None, None), th :: _) =>
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   163
          NetRules.may_unify (#2 (get_cases ctxt))
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   164
            (Logic.strip_assums_concl (#prop (Thm.rep_thm th)))
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   165
          |> map #2
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   166
      | ((Some t, None), _) =>
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   167
          let val name = type_name t in
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   168
            (case lookup_casesT ctxt name of
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   169
              None => error ("No cases rule for type: " ^ quote name)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   170
            | Some thm => [inst_rule t thm])
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   171
          end
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   172
      | ((None, Some thm), _) => [thm]
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   173
      | ((Some t, Some thm), _) => [inst_rule t thm]);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   174
  in Method.rule_tac thms facts end;
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   175
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   176
val cases_meth = Method.METHOD o (FINDGOAL oo cases_tac);
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   177
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   178
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   179
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   180
(** induct method **)
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   181
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   182
(*
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   183
  rule selection:
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   184
        induct         - mathematical induction
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   185
  <x:A> induct         - set induction
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   186
  ...   induct x       - datatype induction
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   187
  ...   induct ... r   - explicit rule
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   188
*)
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   189
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   190
fun induct_tac (ctxt, args) facts =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   191
  let
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   192
    val sg = ProofContext.sign_of ctxt;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   193
    val cert = Thm.cterm_of sg;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   194
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   195
    fun prep_inst (concl, ts) =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   196
      let val xs = vars_of concl; val n = length xs - length ts in
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   197
        if n < 0 then error "More arguments given than in induction rule"
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   198
        else map cert (Library.drop (n, xs)) ~~ map cert ts
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   199
      end;
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   200
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   201
    fun inst_rule insts thm =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   202
      Drule.cterm_instantiate (flat (map2 prep_inst
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   203
        (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of thm)), insts))) thm;
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   204
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   205
    val thms =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   206
      (case (args, facts) of
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   207
        (([], None), []) => [nat_induct]
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   208
      | (([], None), th :: _) =>
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   209
          NetRules.may_unify (#2 (get_induct ctxt))
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   210
            (Logic.strip_assums_concl (#prop (Thm.rep_thm th)))
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   211
          |> map #2
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   212
      | ((insts, None), _) =>
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   213
          let val name = type_name (last_elem (hd insts)) in
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   214
            (case lookup_inductT ctxt name of
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   215
              None => error ("No induct rule for type: " ^ quote name)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   216
            | Some thm => [inst_rule insts thm])
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   217
          end
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   218
      | (([], Some thm), _) => [thm]
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   219
      | ((insts, Some thm), _) => [inst_rule insts thm]);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   220
  in Method.rule_tac thms facts end;
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   221
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   222
val induct_meth = Method.METHOD o (FINDGOAL oo induct_tac);
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   223
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   224
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   225
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   226
(** concrete syntax **)
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   227
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   228
val casesN = "cases";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   229
val inductN = "induct";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   230
val typeN = "type";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   231
val setN = "set";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   232
val ruleN = "rule";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   233
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   234
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   235
(* attributes *)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   236
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   237
fun spec k = (Args.$$$ k -- Args.$$$ ":") |-- Args.!!! Args.name;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   238
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   239
fun attrib sign_of add_type add_set = Scan.depend (fn x =>
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   240
  let val sg = sign_of x in
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   241
    spec typeN >> (add_type o Sign.intern_tycon sg) ||
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   242
    spec setN  >> (add_set o Sign.intern_const sg)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   243
  end >> pair x);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   244
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   245
val cases_attr =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   246
  (Attrib.syntax (attrib Theory.sign_of cases_type_global cases_set_global),
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   247
   Attrib.syntax (attrib ProofContext.sign_of cases_type_local cases_set_local));
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   248
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   249
val induct_attr =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   250
  (Attrib.syntax (attrib Theory.sign_of induct_type_global induct_set_global),
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   251
   Attrib.syntax (attrib ProofContext.sign_of induct_type_local induct_set_local));
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   252
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   253
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   254
(* methods *)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   255
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   256
local
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   257
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   258
fun err k get name =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   259
  (case get name of Some x => x
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   260
  | None => error ("No rule for " ^ k ^ " " ^ quote name));
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   261
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   262
fun rule get_type get_set =
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   263
  Scan.depend (fn ctxt =>
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   264
    let val sg = ProofContext.sign_of ctxt in
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   265
      spec typeN >> (err typeN (get_type ctxt) o Sign.intern_tycon sg) ||
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   266
      spec setN >> (err setN (get_set ctxt) o Sign.intern_const sg)
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   267
    end >> pair ctxt) ||
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   268
  Scan.lift (Args.$$$ ruleN -- Args.$$$ ":") |-- Attrib.local_thm;
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   269
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   270
val cases_rule = rule lookup_casesT lookup_casesS;
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   271
val induct_rule = rule lookup_inductT lookup_inductS;
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   272
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   273
val kind = (Args.$$$ typeN || Args.$$$ setN || Args.$$$ ruleN) -- Args.$$$ ":";
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   274
val term = Scan.unless (Scan.lift kind) Args.local_term;
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
   275
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   276
in
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   277
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   278
val cases_args = Method.syntax (Scan.option term -- Scan.option cases_rule);
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   279
val induct_args = Method.syntax (Args.and_list (Scan.repeat1 term) -- Scan.option induct_rule);
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   280
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   281
end;
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
   282
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
   283
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   284
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   285
(** theory setup **)
6446
583add9799c3 may specify induction predicates as well;
wenzelm
parents: 6442
diff changeset
   286
8278
5928c72b7057 induct: tuned syntax;
wenzelm
parents: 8154
diff changeset
   287
val setup =
8308
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   288
  [GlobalInduct.init, LocalInduct.init,
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   289
   Attrib.add_attributes
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   290
    [(casesN, cases_attr, "cases rule for type or set"),
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   291
     (inductN, induct_attr, "induction rule for type or set")],
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   292
   Method.add_methods
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   293
    [("cases", cases_meth oo cases_args, "case analysis on types or sets"),
45e11d3ccbe4 cases/induct attributes;
wenzelm
parents: 8295
diff changeset
   294
     ("induct", induct_meth oo induct_args, "induction on types or sets")]];
6442
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   295
6a95ceaa977e Proof by induction on types / set / functions.
wenzelm
parents:
diff changeset
   296
end;