| author | nipkow | 
| Tue, 20 Feb 2001 11:27:04 +0100 | |
| changeset 11160 | e0ab13bec5c8 | 
| parent 11035 | bad7568e76e0 | 
| permissions | -rw-r--r-- | 
| 6442 | 1 | (* Title: HOL/Tools/induct_method.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 9230 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 6442 | 5 | |
| 8376 | 6 | Proof by cases and induction on types and sets. | 
| 6442 | 7 | *) | 
| 8 | ||
| 9 | signature INDUCT_METHOD = | |
| 10 | sig | |
| 8431 | 11 | val vars_of: term -> term list | 
| 8695 | 12 | val concls_of: thm -> term list | 
| 9299 | 13 | val simp_case_tac: bool -> simpset -> int -> tactic | 
| 6442 | 14 | val setup: (theory -> theory) list | 
| 15 | end; | |
| 16 | ||
| 17 | structure InductMethod: INDUCT_METHOD = | |
| 18 | struct | |
| 19 | ||
| 8337 | 20 | |
| 10409 | 21 | (** theory context references **) | 
| 22 | ||
| 23 | val inductive_atomize = thms "inductive_atomize"; | |
| 10439 
be2dc95dfe98
inductive_rulify2 accomodates malformed induction rules;
 wenzelm parents: 
10409diff
changeset | 24 | val inductive_rulify1 = thms "inductive_rulify1"; | 
| 
be2dc95dfe98
inductive_rulify2 accomodates malformed induction rules;
 wenzelm parents: 
10409diff
changeset | 25 | val inductive_rulify2 = thms "inductive_rulify2"; | 
| 10409 | 26 | |
| 27 | ||
| 28 | ||
| 8308 | 29 | (** misc utils **) | 
| 8278 | 30 | |
| 9597 | 31 | (* align lists *) | 
| 32 | ||
| 33 | fun align_left msg xs ys = | |
| 34 | let val m = length xs and n = length ys | |
| 35 | in if m < n then error msg else (Library.take (n, xs) ~~ ys) end; | |
| 36 | ||
| 37 | fun align_right msg xs ys = | |
| 38 | let val m = length xs and n = length ys | |
| 39 | in if m < n then error msg else (Library.drop (m - n, xs) ~~ ys) end; | |
| 40 | ||
| 41 | ||
| 8695 | 42 | (* thms and terms *) | 
| 43 | ||
| 10803 | 44 | fun imp_concl_of t = imp_concl_of (#2 (HOLogic.dest_imp t)) handle TERM _ => t; | 
| 45 | val concls_of = map imp_concl_of o HOLogic.dest_conj o HOLogic.dest_Trueprop o Thm.concl_of; | |
| 8344 | 46 | |
| 8278 | 47 | fun vars_of tm = (*ordered left-to-right, preferring right!*) | 
| 8308 | 48 | Term.foldl_aterms (fn (ts, t as Var _) => t :: ts | (ts, _) => ts) ([], tm) | 
| 8278 | 49 | |> Library.distinct |> rev; | 
| 50 | ||
| 8308 | 51 | fun type_name t = | 
| 52 | #1 (Term.dest_Type (Term.type_of t)) | |
| 9597 | 53 |     handle TYPE _ => raise TERM ("Type of term argument is too general", [t]);
 | 
| 54 | ||
| 10409 | 55 | fun prep_inst align cert f (tm, ts) = | 
| 9597 | 56 | let | 
| 10409 | 57 | fun prep_var (x, Some t) = | 
| 58 | let | |
| 59 | val cx = cert x; | |
| 60 |             val {T = xT, sign, ...} = Thm.rep_cterm cx;
 | |
| 61 | val orig_ct = cert t; | |
| 62 | val ct = f orig_ct; | |
| 63 | in | |
| 64 | if Sign.typ_instance sign (#T (Thm.rep_cterm ct), xT) then Some (cx, ct) | |
| 65 | else error (Pretty.string_of (Pretty.block | |
| 66 | [Pretty.str "Ill-typed instantiation:", Pretty.fbrk, | |
| 67 | Display.pretty_cterm orig_ct, Pretty.str " ::", Pretty.brk 1, | |
| 68 | Display.pretty_ctyp (#T (Thm.crep_cterm orig_ct))])) | |
| 69 | end | |
| 9597 | 70 | | prep_var (_, None) = None; | 
| 71 | in | |
| 72 | align "Rule has fewer variables than instantiations given" (vars_of tm) ts | |
| 73 | |> mapfilter prep_var | |
| 74 | end; | |
| 8278 | 75 | |
| 10728 | 76 | |
| 8278 | 77 | |
| 8337 | 78 | (* simplifying cases rules *) | 
| 79 | ||
| 80 | local | |
| 81 | ||
| 82 | (*delete needless equality assumptions*) | |
| 10409 | 83 | val refl_thin = prove_goal HOL.thy "!!P. a = a ==> P ==> P" (fn _ => [assume_tac 1]); | 
| 8337 | 84 | val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE, Pair_inject]; | 
| 85 | val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls; | |
| 86 | ||
| 87 | in | |
| 88 | ||
| 9299 | 89 | fun simp_case_tac solved ss i = | 
| 90 | EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i | |
| 91 | THEN_MAYBE (if solved then no_tac else all_tac); | |
| 8337 | 92 | |
| 93 | end; | |
| 94 | ||
| 95 | ||
| 10542 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 96 | (* resolution and cases *) | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 97 | |
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 98 | local | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 99 | |
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 100 | fun gen_resolveq_tac tac rules i st = | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 101 | Seq.flat (Seq.map (fn rule => tac rule i st) rules); | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 102 | |
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 103 | in | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 104 | |
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 105 | fun resolveq_cases_tac make tac = gen_resolveq_tac (fn (rule, (cases, facts)) => fn i => fn st => | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 106 | Seq.map (rpair (make rule cases)) | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 107 | ((Method.insert_tac facts THEN' tac THEN' Tactic.rtac rule) i st)); | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 108 | |
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 109 | end; | 
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 110 | |
| 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 111 | |
| 8278 | 112 | |
| 113 | (** cases method **) | |
| 114 | ||
| 8308 | 115 | (* | 
| 116 | rule selection: | |
| 117 | cases - classical case split | |
| 8376 | 118 | cases t - datatype exhaustion | 
| 119 | <x:A> cases ... - set elimination | |
| 8451 | 120 | ... cases ... R - explicit rule | 
| 8308 | 121 | *) | 
| 8278 | 122 | |
| 9066 | 123 | val case_split = RuleCases.name ["True", "False"] case_split_thm; | 
| 124 | ||
| 8344 | 125 | local | 
| 126 | ||
| 9299 | 127 | fun simplified_cases ctxt cases thm = | 
| 128 | let | |
| 129 | val nprems = Thm.nprems_of thm; | |
| 130 | val opt_cases = | |
| 131 | Library.replicate (nprems - Int.min (nprems, length cases)) None @ | |
| 132 | map Some (Library.take (nprems, cases)); | |
| 8337 | 133 | |
| 9299 | 134 | val tac = simp_case_tac true (Simplifier.get_local_simpset ctxt); | 
| 135 | fun simp ((i, c), (th, cs)) = | |
| 136 | (case try (Tactic.rule_by_tactic (tac i)) th of | |
| 9313 | 137 | None => (th, c :: cs) | 
| 138 | | Some th' => (th', None :: cs)); | |
| 9299 | 139 | |
| 140 | val (thm', opt_cases') = foldr simp (1 upto Thm.nprems_of thm ~~ opt_cases, (thm, [])); | |
| 141 | in (thm', mapfilter I opt_cases') end; | |
| 142 | ||
| 9624 
de254f375477
changed 'opaque' option to 'open' (opaque is default);
 wenzelm parents: 
9597diff
changeset | 143 | fun cases_tac (ctxt, ((simplified, open_parms), args)) facts = | 
| 8308 | 144 | let | 
| 145 | val sg = ProofContext.sign_of ctxt; | |
| 146 | val cert = Thm.cterm_of sg; | |
| 8278 | 147 | |
| 9597 | 148 | fun inst_rule insts thm = | 
| 9914 | 149 | (align_left "Rule has fewer premises than arguments given" (Thm.prems_of thm) insts | 
| 10409 | 150 | |> (flat o map (prep_inst align_left cert I)) | 
| 9597 | 151 | |> Drule.cterm_instantiate) thm; | 
| 6442 | 152 | |
| 8376 | 153 | fun find_cases th = | 
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 154 | NetRules.may_unify (#2 (InductAttrib.get_cases ctxt)) | 
| 8376 | 155 | (Logic.strip_assums_concl (#prop (Thm.rep_thm th))); | 
| 156 | ||
| 157 | val rules = | |
| 10803 | 158 | (case (fst args, facts) of | 
| 9597 | 159 | (([], None), []) => [RuleCases.add case_split] | 
| 160 | | ((insts, None), []) => | |
| 161 | let | |
| 162 | val name = type_name (hd (flat (map (mapfilter I) insts))) | |
| 163 | handle Library.LIST _ => error "Unable to figure out type cases rule" | |
| 164 | in | |
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 165 | (case InductAttrib.lookup_casesT ctxt name of | 
| 8308 | 166 |               None => error ("No cases rule for type: " ^ quote name)
 | 
| 9597 | 167 | | Some thm => [(inst_rule insts thm, RuleCases.get thm)]) | 
| 8308 | 168 | end | 
| 9597 | 169 | | (([], None), th :: _) => map (RuleCases.add o #2) (find_cases th) | 
| 170 | | ((insts, None), th :: _) => | |
| 9299 | 171 | (case find_cases th of (*may instantiate first rule only!*) | 
| 9597 | 172 | (_, thm) :: _ => [(inst_rule insts thm, RuleCases.get thm)] | 
| 8376 | 173 | | [] => []) | 
| 9597 | 174 | | (([], Some thm), _) => [RuleCases.add thm] | 
| 10803 | 175 | | ((insts, Some thm), _) => [(inst_rule insts thm, RuleCases.get thm)]) | 
| 176 | |> map (Library.apfst (Attrib.read_inst' (snd args) ctxt)); | |
| 8376 | 177 | |
| 9299 | 178 | val cond_simp = if simplified then simplified_cases ctxt else rpair; | 
| 179 | ||
| 10527 | 180 | fun prep_rule (thm, (cases, n)) = Seq.map (apsnd (rpair (drop (n, facts))) o cond_simp cases) | 
| 181 | (Method.multi_resolves (take (n, facts)) [thm]); | |
| 10409 | 182 | in | 
| 10542 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 183 | resolveq_cases_tac (RuleCases.make open_parms) (K all_tac) | 
| 10409 | 184 | (Seq.flat (Seq.map prep_rule (Seq.of_list rules))) | 
| 185 | end; | |
| 8278 | 186 | |
| 8344 | 187 | in | 
| 188 | ||
| 8671 | 189 | val cases_meth = Method.METHOD_CASES o (HEADGOAL oo cases_tac); | 
| 8278 | 190 | |
| 8344 | 191 | end; | 
| 192 | ||
| 8278 | 193 | |
| 194 | ||
| 195 | (** induct method **) | |
| 196 | ||
| 8308 | 197 | (* | 
| 198 | rule selection: | |
| 8376 | 199 | induct x - datatype induction | 
| 200 | <x:A> induct ... - set induction | |
| 8451 | 201 | ... induct ... R - explicit rule | 
| 8308 | 202 | *) | 
| 8278 | 203 | |
| 8344 | 204 | local | 
| 205 | ||
| 11035 | 206 | val atomize_cterm = Thm.cterm_fun AutoBind.drop_judgment o hol_rewrite_cterm inductive_atomize; | 
| 10439 
be2dc95dfe98
inductive_rulify2 accomodates malformed induction rules;
 wenzelm parents: 
10409diff
changeset | 207 | val atomize_tac = Tactic.rewrite_goal_tac inductive_atomize; | 
| 11035 | 208 | val rulify_cterm = hol_rewrite_cterm inductive_rulify2 o hol_rewrite_cterm inductive_rulify1; | 
| 10439 
be2dc95dfe98
inductive_rulify2 accomodates malformed induction rules;
 wenzelm parents: 
10409diff
changeset | 209 | |
| 
be2dc95dfe98
inductive_rulify2 accomodates malformed induction rules;
 wenzelm parents: 
10409diff
changeset | 210 | val rulify_tac = | 
| 
be2dc95dfe98
inductive_rulify2 accomodates malformed induction rules;
 wenzelm parents: 
10409diff
changeset | 211 | Tactic.rewrite_goal_tac inductive_rulify1 THEN' | 
| 
be2dc95dfe98
inductive_rulify2 accomodates malformed induction rules;
 wenzelm parents: 
10409diff
changeset | 212 | Tactic.rewrite_goal_tac inductive_rulify2 THEN' | 
| 10803 | 213 | Tactic.norm_hhf_tac; | 
| 10409 | 214 | |
| 215 | fun rulify_cases cert = | |
| 10803 | 216 | let | 
| 217 | val ruly = Thm.term_of o rulify_cterm o cert; | |
| 218 |     fun ruly_case {fixes, assumes, binds} =
 | |
| 10814 | 219 |       {fixes = fixes, assumes = map ruly assumes,
 | 
| 220 | binds = map (apsnd (apsome (AutoBind.drop_judgment o ruly))) binds}; | |
| 10871 | 221 | in map (apsnd ruly_case) ooo RuleCases.make_raw end; | 
| 10409 | 222 | |
| 10455 | 223 | val weak_strip_tac = REPEAT o Tactic.match_tac [impI, allI, ballI]; | 
| 224 | ||
| 10409 | 225 | |
| 8376 | 226 | infix 1 THEN_ALL_NEW_CASES; | 
| 227 | ||
| 228 | fun (tac1 THEN_ALL_NEW_CASES tac2) i st = | |
| 229 | st |> Seq.THEN (tac1 i, (fn (st', cases) => | |
| 8540 | 230 | Seq.map (rpair cases) (Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st) st'))); | 
| 8376 | 231 | |
| 232 | ||
| 8330 | 233 | fun induct_rule ctxt t = | 
| 234 | let val name = type_name t in | |
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 235 | (case InductAttrib.lookup_inductT ctxt name of | 
| 8330 | 236 |       None => error ("No induct rule for type: " ^ quote name)
 | 
| 8332 | 237 | | Some thm => (name, thm)) | 
| 8330 | 238 | end; | 
| 239 | ||
| 8332 | 240 | fun join_rules [(_, thm)] = thm | 
| 8330 | 241 | | join_rules raw_thms = | 
| 242 | let | |
| 8332 | 243 | val thms = (map (apsnd Drule.freeze_all) raw_thms); | 
| 244 | fun eq_prems ((_, th1), (_, th2)) = | |
| 245 | Term.aconvs (Thm.prems_of th1, Thm.prems_of th2); | |
| 8330 | 246 | in | 
| 8332 | 247 | (case Library.gen_distinct eq_prems thms of | 
| 248 | [(_, thm)] => | |
| 249 | let | |
| 250 | val cprems = Drule.cprems_of thm; | |
| 251 | val asms = map Thm.assume cprems; | |
| 252 | fun strip (_, th) = Drule.implies_elim_list th asms; | |
| 253 | in | |
| 254 | foldr1 (fn (th, th') => [th, th'] MRS conjI) (map strip thms) | |
| 255 | |> Drule.implies_intr_list cprems | |
| 256 | |> Drule.standard | |
| 257 | end | |
| 258 | | [] => error "No rule given" | |
| 259 |         | bads => error ("Incompatible rules for " ^ commas_quote (map #1 bads)))
 | |
| 8330 | 260 | end; | 
| 261 | ||
| 8376 | 262 | |
| 9624 
de254f375477
changed 'opaque' option to 'open' (opaque is default);
 wenzelm parents: 
9597diff
changeset | 263 | fun induct_tac (ctxt, ((stripped, open_parms), args)) facts = | 
| 8308 | 264 | let | 
| 265 | val sg = ProofContext.sign_of ctxt; | |
| 266 | val cert = Thm.cterm_of sg; | |
| 267 | ||
| 268 | fun inst_rule insts thm = | |
| 9597 | 269 | (align_right "Rule has fewer conclusions than arguments given" (concls_of thm) insts | 
| 10409 | 270 | |> (flat o map (prep_inst align_right cert atomize_cterm)) | 
| 9597 | 271 | |> Drule.cterm_instantiate) thm; | 
| 8278 | 272 | |
| 8376 | 273 | fun find_induct th = | 
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 274 | NetRules.may_unify (#2 (InductAttrib.get_induct ctxt)) | 
| 8376 | 275 | (Logic.strip_assums_concl (#prop (Thm.rep_thm th))); | 
| 276 | ||
| 277 | val rules = | |
| 10803 | 278 | (case (fst args, facts) of | 
| 8540 | 279 | (([], None), []) => [] | 
| 8376 | 280 | | ((insts, None), []) => | 
| 8695 | 281 | let val thms = map (induct_rule ctxt o last_elem o mapfilter I) insts | 
| 282 | handle Library.LIST _ => error "Unable to figure out type induction rule" | |
| 8376 | 283 | in [(inst_rule insts (join_rules thms), RuleCases.get (#2 (hd thms)))] end | 
| 284 | | (([], None), th :: _) => map (RuleCases.add o #2) (find_induct th) | |
| 285 | | ((insts, None), th :: _) => | |
| 9299 | 286 | (case find_induct th of (*may instantiate first rule only!*) | 
| 287 | (_, thm) :: _ => [(inst_rule insts thm, RuleCases.get thm)] | |
| 8376 | 288 | | [] => []) | 
| 289 | | (([], Some thm), _) => [RuleCases.add thm] | |
| 10803 | 290 | | ((insts, Some thm), _) => [(inst_rule insts thm, RuleCases.get thm)]) | 
| 291 | |> map (Library.apfst (Attrib.read_inst' (snd args) ctxt)); | |
| 8376 | 292 | |
| 10527 | 293 | fun prep_rule (thm, (cases, n)) = | 
| 294 | Seq.map (rpair (cases, drop (n, facts))) (Method.multi_resolves (take (n, facts)) [thm]); | |
| 10542 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 295 | val tac = resolveq_cases_tac (rulify_cases cert open_parms) atomize_tac | 
| 9624 
de254f375477
changed 'opaque' option to 'open' (opaque is default);
 wenzelm parents: 
9597diff
changeset | 296 | (Seq.flat (Seq.map prep_rule (Seq.of_list rules))); | 
| 8344 | 297 | in | 
| 10542 
92cd56dfc17e
resolveq_cases_tac moved here from Pure/Isar/method.ML;
 wenzelm parents: 
10527diff
changeset | 298 | tac THEN_ALL_NEW_CASES (rulify_tac THEN' (if stripped then weak_strip_tac else K all_tac)) | 
| 8344 | 299 | end; | 
| 300 | ||
| 301 | in | |
| 8278 | 302 | |
| 8671 | 303 | val induct_meth = Method.METHOD_CASES o (HEADGOAL oo induct_tac); | 
| 8278 | 304 | |
| 8344 | 305 | end; | 
| 306 | ||
| 8278 | 307 | |
| 308 | ||
| 309 | (** concrete syntax **) | |
| 310 | ||
| 8337 | 311 | val simplifiedN = "simplified"; | 
| 8344 | 312 | val strippedN = "stripped"; | 
| 9624 
de254f375477
changed 'opaque' option to 'open' (opaque is default);
 wenzelm parents: 
9597diff
changeset | 313 | val openN = "open"; | 
| 8308 | 314 | val ruleN = "rule"; | 
| 10803 | 315 | val ofN = "of"; | 
| 8308 | 316 | |
| 8278 | 317 | local | 
| 6442 | 318 | |
| 8308 | 319 | fun err k get name = | 
| 320 | (case get name of Some x => x | |
| 321 |   | None => error ("No rule for " ^ k ^ " " ^ quote name));
 | |
| 6442 | 322 | |
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 323 | fun spec k = (Args.$$$ k -- Args.colon) |-- Args.!!! Args.name; | 
| 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 324 | |
| 8308 | 325 | fun rule get_type get_set = | 
| 326 | Scan.depend (fn ctxt => | |
| 327 | let val sg = ProofContext.sign_of ctxt in | |
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 328 | spec InductAttrib.typeN >> (err InductAttrib.typeN (get_type ctxt) o Sign.intern_tycon sg) || | 
| 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 329 | spec InductAttrib.setN >> (err InductAttrib.setN (get_set ctxt) o Sign.intern_const sg) | 
| 8308 | 330 | end >> pair ctxt) || | 
| 8815 | 331 | Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.local_thm; | 
| 6442 | 332 | |
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 333 | val cases_rule = rule InductAttrib.lookup_casesT InductAttrib.lookup_casesS; | 
| 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 334 | val induct_rule = rule InductAttrib.lookup_inductT InductAttrib.lookup_inductS; | 
| 6442 | 335 | |
| 10803 | 336 | val kind_inst = | 
| 337 | (Args.$$$ InductAttrib.typeN || Args.$$$ InductAttrib.setN || Args.$$$ ruleN || Args.$$$ ofN) | |
| 338 | -- Args.colon; | |
| 339 | val term = Scan.unless (Scan.lift kind_inst) Args.local_term; | |
| 340 | val term_dummy = Scan.unless (Scan.lift kind_inst) | |
| 8695 | 341 | (Scan.lift (Args.$$$ "_") >> K None || Args.local_term >> Some); | 
| 6446 | 342 | |
| 9597 | 343 | val instss = Args.and_list (Scan.repeat1 term_dummy); | 
| 344 | ||
| 10803 | 345 | (* FIXME Attrib.insts': better use actual term args *) | 
| 346 | val rule_insts = | |
| 347 | Scan.lift (Scan.optional ((Args.$$$ ofN -- Args.colon) |-- Args.!!! Attrib.insts') ([], [])); | |
| 348 | ||
| 8278 | 349 | in | 
| 350 | ||
| 9299 | 351 | val cases_args = Method.syntax | 
| 10803 | 352 | (Args.mode simplifiedN -- Args.mode openN -- (instss -- Scan.option cases_rule -- rule_insts)); | 
| 9299 | 353 | |
| 8695 | 354 | val induct_args = Method.syntax | 
| 10803 | 355 | (Args.mode strippedN -- Args.mode openN -- (instss -- Scan.option induct_rule -- rule_insts)); | 
| 8278 | 356 | |
| 357 | end; | |
| 6446 | 358 | |
| 359 | ||
| 6442 | 360 | |
| 8278 | 361 | (** theory setup **) | 
| 6446 | 362 | |
| 8278 | 363 | val setup = | 
| 10271 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 364 | [Method.add_methods | 
| 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 365 | [(InductAttrib.casesN, cases_meth oo cases_args, "case analysis on types or sets"), | 
| 
45b996639c45
split over two files: induct_attrib.ML, induct_method.ML;
 wenzelm parents: 
10013diff
changeset | 366 | (InductAttrib.inductN, induct_meth oo induct_args, "induction on types or sets")], | 
| 9066 | 367 |    (#1 o PureThy.add_thms [(("case_split", case_split), [])])];
 | 
| 6442 | 368 | |
| 369 | end; |