| author | wenzelm | 
| Sat, 29 Dec 2001 18:36:12 +0100 | |
| changeset 12610 | 8b9845807f77 | 
| parent 12381 | 5177845a34f5 | 
| child 13105 | 3d1e7a199bdc | 
| permissions | -rw-r--r-- | 
| 11658 | 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 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 5 | |
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 6 | Declaration of rules for cases and induction. | 
| 
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 | |
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 9 | signature INDUCT_ATTRIB = | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 10 | sig | 
| 11730 | 11 | val vars_of: term -> term list | 
| 11656 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 12 | val dest_global_rules: theory -> | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 13 |     {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 | 14 | 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 | 15 | val print_global_rules: theory -> unit | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 16 | val dest_local_rules: Proof.context -> | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 17 |     {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 | 18 | 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 | 19 | val print_local_rules: Proof.context -> unit | 
| 11730 | 20 | 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 | 21 | val lookup_casesS : Proof.context -> string -> thm option | 
| 11730 | 22 | 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 | 23 | val lookup_inductS : Proof.context -> string -> thm option | 
| 11730 | 24 | val find_casesT: Proof.context -> typ -> thm list | 
| 25 | val find_casesS: Proof.context -> thm -> thm list | |
| 26 | val find_inductT: Proof.context -> typ -> thm list | |
| 27 | 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 | 28 | val cases_type_global: string -> theory attribute | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 29 | val cases_set_global: string -> theory attribute | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 30 | 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 | 31 | 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 | 32 | val induct_type_global: string -> theory attribute | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 33 | val induct_set_global: string -> theory attribute | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 34 | 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 | 35 | 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 | 36 | val casesN: string | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 37 | val inductN: string | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 38 | val typeN: string | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 39 | val setN: string | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 40 | val setup: (theory -> theory) list | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 41 | end; | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 42 | |
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 43 | structure InductAttrib: INDUCT_ATTRIB = | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 44 | struct | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 45 | |
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 46 | |
| 11730 | 47 | (** misc utils **) | 
| 48 | ||
| 49 | (* encode_type -- for indexing purposes *) | |
| 50 | ||
| 51 | fun encode_type (Type (c, Ts)) = Term.list_comb (Const (c, dummyT), map encode_type Ts) | |
| 52 | | encode_type (TFree (a, _)) = Free (a, dummyT) | |
| 53 | | encode_type (TVar (a, _)) = Var (a, dummyT); | |
| 54 | ||
| 55 | ||
| 56 | (* variables -- ordered left-to-right, preferring right *) | |
| 57 | ||
| 58 | local | |
| 59 | ||
| 60 | fun rev_vars_of tm = | |
| 61 | Term.foldl_aterms (fn (ts, t as Var _) => t :: ts | (ts, _) => ts) ([], tm) | |
| 62 | |> Library.distinct; | |
| 63 | ||
| 64 | val mk_var = encode_type o #2 o Term.dest_Var; | |
| 65 | ||
| 66 | in | |
| 67 | ||
| 68 | val vars_of = rev o rev_vars_of; | |
| 69 | ||
| 70 | fun first_var thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle LIST _ => | |
| 71 |   raise THM ("No variables in first premise of rule", 0, [thm]);
 | |
| 72 | ||
| 73 | fun last_var thm = mk_var (hd (rev_vars_of (Thm.concl_of thm))) handle LIST _ => | |
| 74 |   raise THM ("No variables in conclusion of rule", 0, [thm]);
 | |
| 75 | ||
| 76 | end; | |
| 77 | ||
| 78 | ||
| 79 | ||
| 11656 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 80 | (** global and local induct data **) | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 81 | |
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 82 | (* rules *) | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 83 | |
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 84 | type rules = (string * thm) NetRules.T; | 
| 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 85 | |
| 12381 | 86 | val init_rules = | 
| 87 | NetRules.init (fn ((s1: string, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm (th1, th2)); | |
| 11656 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 88 | |
| 11730 | 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 | 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 | 93 | in Pretty.writeln (Pretty.big_list kind (map (prt x) thms)) end; | 
| 94 | ||
| 95 | fun print_all_rules prt x ((casesT, casesS), (inductT, inductS)) = | |
| 96 | (print_rules prt "induct type:" x inductT; | |
| 97 | print_rules prt "induct set:" x inductS; | |
| 98 | print_rules prt "cases type:" x casesT; | |
| 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 | 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 | 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 | 109 | val empty = | 
| 110 | ((init_rules (first_var o #2), init_rules (Thm.major_prem_of o #2)), | |
| 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 | 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 | 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 | 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 | 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 | 151 | val lookup_casesT = lookup_rule o #1 o #1 o LocalInduct.get; | 
| 152 | val lookup_casesS = lookup_rule o #2 o #1 o LocalInduct.get; | |
| 153 | val lookup_inductT = lookup_rule o #1 o #2 o LocalInduct.get; | |
| 154 | val lookup_inductS = lookup_rule o #2 o #2 o LocalInduct.get; | |
| 155 | ||
| 11656 
e499dceca569
moved HOL/Tools/induct_attrib.ML to Pure/Isar/induct_attrib.ML;
 wenzelm parents: diff
changeset | 156 | |
| 11730 | 157 | fun find_rules which how ctxt x = | 
| 11784 | 158 | map snd (NetRules.retrieve (which (LocalInduct.get ctxt)) (how x)); | 
| 11730 | 159 | |
| 160 | val find_casesT = find_rules (#1 o #1) encode_type; | |
| 161 | val find_casesS = find_rules (#2 o #1) Thm.concl_of; | |
| 162 | val find_inductT = find_rules (#1 o #2) encode_type; | |
| 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 | 207 | fun spec k cert = | 
| 208 | (Args.$$$ k -- Args.colon) |-- Args.!!! (Args.name >> cert) || | |
| 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 | 213 | spec typeN (Sign.certify_tyname sg o Sign.intern_tycon sg) >> add_type || | 
| 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; |