src/Pure/Isar/induct_attrib.ML
author kleing
Mon, 21 Jun 2004 10:25:57 +0200
changeset 14981 e73f8140af78
parent 13105 3d1e7a199bdc
child 15570 8d8c70b41bab
permissions -rw-r--r--
Merged in license change from Isabelle2004
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11658
4200394242c5 Isar/induct_attrib.ML;
wenzelm
parents: 11656
diff changeset
     1
(*  Title:      Pure/Isar/induct_attrib.ML
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     4
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     5
Declaration of rules for cases and induction.
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     6
*)
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     7
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     8
signature INDUCT_ATTRIB =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
     9
sig
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    10
  val vars_of: term -> term list
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    11
  val dest_global_rules: theory ->
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    12
    {type_cases: (string * thm) list, set_cases: (string * thm) list,
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    13
      type_induct: (string * thm) list, set_induct: (string * thm) list}
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    14
  val print_global_rules: theory -> unit
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    15
  val dest_local_rules: Proof.context ->
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    16
    {type_cases: (string * thm) list, set_cases: (string * thm) list,
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    17
      type_induct: (string * thm) list, set_induct: (string * thm) list}
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    18
  val print_local_rules: Proof.context -> unit
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    19
  val lookup_casesT : Proof.context -> string -> thm option
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    20
  val lookup_casesS : Proof.context -> string -> thm option
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    21
  val lookup_inductT : Proof.context -> string -> thm option
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    22
  val lookup_inductS : Proof.context -> string -> thm option
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    23
  val find_casesT: Proof.context -> typ -> thm list
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    24
  val find_casesS: Proof.context -> thm -> thm list
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    25
  val find_inductT: Proof.context -> typ -> thm list
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    26
  val find_inductS: Proof.context -> thm -> thm list
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    27
  val cases_type_global: string -> theory attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    28
  val cases_set_global: string -> theory attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    29
  val cases_type_local: string -> Proof.context attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    30
  val cases_set_local: string -> Proof.context attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    31
  val induct_type_global: string -> theory attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    32
  val induct_set_global: string -> theory attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    33
  val induct_type_local: string -> Proof.context attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    34
  val induct_set_local: string -> Proof.context attribute
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    35
  val casesN: string
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    36
  val inductN: string
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    37
  val typeN: string
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    38
  val setN: string
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    39
  val setup: (theory -> theory) list
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    40
end;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    41
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    42
structure InductAttrib: INDUCT_ATTRIB =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    43
struct
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    44
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    45
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    46
(** misc utils **)
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    47
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    48
(* encode_type -- for indexing purposes *)
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    49
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    50
fun encode_type (Type (c, Ts)) = Term.list_comb (Const (c, dummyT), map encode_type Ts)
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    51
  | encode_type (TFree (a, _)) = Free (a, dummyT)
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    52
  | encode_type (TVar (a, _)) = Var (a, dummyT);
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    53
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    54
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    55
(* variables -- ordered left-to-right, preferring right *)
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    56
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    57
local
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    58
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    59
fun rev_vars_of tm =
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    60
  Term.foldl_aterms (fn (ts, t as Var _) => t :: ts | (ts, _) => ts) ([], tm)
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    61
  |> Library.distinct;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    62
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    63
val mk_var = encode_type o #2 o Term.dest_Var;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    64
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    65
in
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    66
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    67
val vars_of = rev o rev_vars_of;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    68
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    69
fun first_var thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle LIST _ =>
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    70
  raise THM ("No variables in first premise of rule", 0, [thm]);
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    71
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    72
fun last_var thm = mk_var (hd (rev_vars_of (Thm.concl_of thm))) handle LIST _ =>
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    73
  raise THM ("No variables in conclusion of rule", 0, [thm]);
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    74
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    75
end;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    76
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    77
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    78
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    79
(** global and local induct data **)
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    80
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    81
(* rules *)
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    82
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    83
type rules = (string * thm) NetRules.T;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    84
12381
5177845a34f5 simplified NetRules;
wenzelm
parents: 12311
diff changeset
    85
val init_rules =
13105
3d1e7a199bdc use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents: 12381
diff changeset
    86
  NetRules.init (fn ((s1: string, th1), (s2, th2)) => s1 = s2 andalso
3d1e7a199bdc use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents: 12381
diff changeset
    87
    Drule.eq_thm_prop (th1, th2));
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    88
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
    89
fun lookup_rule (rs: rules) name = Library.assoc (NetRules.rules rs, name);
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    90
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    91
fun print_rules prt kind x rs =
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
    92
  let val thms = map snd (NetRules.rules rs)
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    93
  in Pretty.writeln (Pretty.big_list kind (map (prt x) thms)) end;
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    94
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    95
fun print_all_rules prt x ((casesT, casesS), (inductT, inductS)) =
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    96
 (print_rules prt "induct type:" x inductT;
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    97
  print_rules prt "induct set:" x inductS;
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    98
  print_rules prt "cases type:" x casesT;
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 11791
diff changeset
    99
  print_rules prt "cases set:" x casesS);
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   100
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   101
11658
4200394242c5 Isar/induct_attrib.ML;
wenzelm
parents: 11656
diff changeset
   102
(* theory data kind 'Isar/induction' *)
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   103
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   104
structure GlobalInductArgs =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   105
struct
11658
4200394242c5 Isar/induct_attrib.ML;
wenzelm
parents: 11656
diff changeset
   106
  val name = "Isar/induction";
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   107
  type T = (rules * rules) * (rules * rules);
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   108
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   109
  val empty =
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   110
    ((init_rules (first_var o #2), init_rules (Thm.major_prem_of o #2)),
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   111
     (init_rules (last_var o #2), init_rules (Thm.major_prem_of o #2)));
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   112
  val copy = I;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   113
  val prep_ext = I;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   114
  fun merge (((casesT1, casesS1), (inductT1, inductS1)),
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   115
      ((casesT2, casesS2), (inductT2, inductS2))) =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   116
    ((NetRules.merge (casesT1, casesT2), NetRules.merge (casesS1, casesS2)),
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   117
      (NetRules.merge (inductT1, inductT2), NetRules.merge (inductS1, inductS2)));
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   118
12102
a2deb1c3cd9b make SML/XL of NJ happy;
wenzelm
parents: 12055
diff changeset
   119
  fun print x = print_all_rules Display.pretty_thm_sg x;
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   120
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   121
  fun dest ((casesT, casesS), (inductT, inductS)) =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   122
    {type_cases = NetRules.rules casesT,
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   123
     set_cases = NetRules.rules casesS,
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   124
     type_induct = NetRules.rules inductT,
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   125
     set_induct = NetRules.rules inductS};
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   126
end;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   127
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   128
structure GlobalInduct = TheoryDataFun(GlobalInductArgs);
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   129
val print_global_rules = GlobalInduct.print;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   130
val dest_global_rules = GlobalInductArgs.dest o GlobalInduct.get;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   131
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   132
11658
4200394242c5 Isar/induct_attrib.ML;
wenzelm
parents: 11656
diff changeset
   133
(* proof data kind 'Isar/induction' *)
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   134
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   135
structure LocalInductArgs =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   136
struct
11658
4200394242c5 Isar/induct_attrib.ML;
wenzelm
parents: 11656
diff changeset
   137
  val name = "Isar/induction";
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   138
  type T = GlobalInductArgs.T;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   139
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   140
  fun init thy = GlobalInduct.get thy;
12102
a2deb1c3cd9b make SML/XL of NJ happy;
wenzelm
parents: 12055
diff changeset
   141
  fun print x = print_all_rules ProofContext.pretty_thm x;
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   142
end;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   143
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   144
structure LocalInduct = ProofDataFun(LocalInductArgs);
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   145
val print_local_rules = LocalInduct.print;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   146
val dest_local_rules = GlobalInductArgs.dest o LocalInduct.get;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   147
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   148
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   149
(* access rules *)
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   150
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   151
val lookup_casesT = lookup_rule o #1 o #1 o LocalInduct.get;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   152
val lookup_casesS = lookup_rule o #2 o #1 o LocalInduct.get;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   153
val lookup_inductT = lookup_rule o #1 o #2 o LocalInduct.get;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   154
val lookup_inductS = lookup_rule o #2 o #2 o LocalInduct.get;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   155
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   156
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   157
fun find_rules which how ctxt x =
11784
b66b198ee29a tuned NetRules;
wenzelm
parents: 11730
diff changeset
   158
  map snd (NetRules.retrieve (which (LocalInduct.get ctxt)) (how x));
11730
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   159
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   160
val find_casesT = find_rules (#1 o #1) encode_type;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   161
val find_casesS = find_rules (#2 o #1) Thm.concl_of;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   162
val find_inductT = find_rules (#1 o #2) encode_type;
418533653668 removed get_cases, get_induct;
wenzelm
parents: 11665
diff changeset
   163
val find_inductS = find_rules (#2 o #2) Thm.concl_of;
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   164
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   165
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   166
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   167
(** attributes **)
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   168
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   169
local
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   170
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   171
fun mk_att f g h name arg =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   172
  let val (x, thm) = h arg in (f (g (name, thm)) x, thm) end;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   173
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   174
fun add_casesT rule x = apfst (apfst (NetRules.insert rule)) x;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   175
fun add_casesS rule x = apfst (apsnd (NetRules.insert rule)) x;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   176
fun add_inductT rule x = apsnd (apfst (NetRules.insert rule)) x;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   177
fun add_inductS rule x = apsnd (apsnd (NetRules.insert rule)) x;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   178
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   179
fun consumes0 x = RuleCases.consumes_default 0 x;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   180
fun consumes1 x = RuleCases.consumes_default 1 x;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   181
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   182
in
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   183
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   184
val cases_type_global = mk_att GlobalInduct.map add_casesT consumes0;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   185
val cases_set_global = mk_att GlobalInduct.map add_casesS consumes1;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   186
val induct_type_global = mk_att GlobalInduct.map add_inductT consumes0;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   187
val induct_set_global = mk_att GlobalInduct.map add_inductS consumes1;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   188
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   189
val cases_type_local = mk_att LocalInduct.map add_casesT consumes0;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   190
val cases_set_local = mk_att LocalInduct.map add_casesS consumes1;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   191
val induct_type_local = mk_att LocalInduct.map add_inductT consumes0;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   192
val induct_set_local = mk_att LocalInduct.map add_inductS consumes1;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   193
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   194
end;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   195
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   196
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   197
(** concrete syntax **)
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   198
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   199
val casesN = "cases";
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   200
val inductN = "induct";
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   201
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   202
val typeN = "type";
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   203
val setN = "set";
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   204
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   205
local
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   206
11791
2c201f3b018e allow empty set/type name;
wenzelm
parents: 11784
diff changeset
   207
fun spec k cert =
2c201f3b018e allow empty set/type name;
wenzelm
parents: 11784
diff changeset
   208
  (Args.$$$ k -- Args.colon) |-- Args.!!! (Args.name >> cert) ||
2c201f3b018e allow empty set/type name;
wenzelm
parents: 11784
diff changeset
   209
  Args.$$$ k >> K "";
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   210
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   211
fun attrib sign_of add_type add_set = Scan.depend (fn x =>
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   212
  let val sg = sign_of x in
11791
2c201f3b018e allow empty set/type name;
wenzelm
parents: 11784
diff changeset
   213
    spec typeN (Sign.certify_tyname sg o Sign.intern_tycon sg) >> add_type ||
2c201f3b018e allow empty set/type name;
wenzelm
parents: 11784
diff changeset
   214
    spec setN (Sign.certify_const sg o Sign.intern_const sg) >> add_set
11656
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   215
  end >> pair x);
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   216
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   217
in
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   218
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   219
val cases_attr =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   220
  (Attrib.syntax (attrib Theory.sign_of cases_type_global cases_set_global),
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   221
   Attrib.syntax (attrib ProofContext.sign_of cases_type_local cases_set_local));
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   222
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   223
val induct_attr =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   224
  (Attrib.syntax (attrib Theory.sign_of induct_type_global induct_set_global),
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   225
   Attrib.syntax (attrib ProofContext.sign_of induct_type_local induct_set_local));
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   226
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   227
end;
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   228
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   229
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   230
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   231
(** theory setup **)
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   232
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   233
val setup =
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   234
  [GlobalInduct.init, LocalInduct.init,
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   235
   Attrib.add_attributes
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   236
    [(casesN, cases_attr, "declaration of cases rule for type or set"),
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   237
     (inductN, induct_attr, "declaration of induction rule for type or set")]];
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   238
e499dceca569 moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
wenzelm
parents:
diff changeset
   239
end;