| author | wenzelm | 
| Thu, 07 Apr 2011 17:38:17 +0200 | |
| changeset 42266 | f87e0be80a3f | 
| parent 40722 | 441260986b63 | 
| child 42361 | 23f352990944 | 
| permissions | -rw-r--r-- | 
| 28308 | 1 | (* Title: Tools/induct_tacs.ML | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 2 | Author: Makarius | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 3 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 4 | Unstructured induction and cases analysis. | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 5 | *) | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 6 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 7 | signature INDUCT_TACS = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 8 | sig | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 9 | val case_tac: Proof.context -> string -> int -> tactic | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 10 | val case_rule_tac: Proof.context -> string -> thm -> int -> tactic | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 11 | val induct_tac: Proof.context -> string option list list -> int -> tactic | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 12 | val induct_rules_tac: Proof.context -> string option list list -> thm list -> int -> tactic | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 13 | val setup: theory -> theory | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 14 | end | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 15 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 16 | structure InductTacs: INDUCT_TACS = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 17 | struct | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 18 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 19 | (* case analysis *) | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 20 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 21 | fun check_type ctxt t = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 22 | let | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 23 | val u = singleton (Variable.polymorphic ctxt) t; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 24 | val U = Term.fastype_of u; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 25 | val _ = Term.is_TVar U andalso | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 26 |       error ("Cannot determine type of " ^ Syntax.string_of_term ctxt u);
 | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 27 | in (u, U) end; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 28 | |
| 30541 | 29 | fun gen_case_tac ctxt0 s opt_rule i st = | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 30 | let | 
| 32863 | 31 | val (_, ctxt) = Variable.focus_subgoal i st ctxt0; | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 32 | val rule = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 33 | (case opt_rule of | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 34 | SOME rule => rule | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 35 | | NONE => | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 36 | (case Induct.find_casesT ctxt | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 37 | (#2 (check_type ctxt (ProofContext.read_term_schematic ctxt s))) of | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 38 | rule :: _ => rule | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 39 |           | [] => @{thm case_split}));
 | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 40 | val _ = Method.trace ctxt [rule]; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 41 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 42 | val xi = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 43 | (case Induct.vars_of (Thm.term_of (Thm.cprem_of rule 1)) of | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 44 | Var (xi, _) :: _ => xi | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 45 |       | _ => raise THM ("Malformed cases rule", 0, [rule]));
 | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 46 | in res_inst_tac ctxt [(xi, s)] rule i st end | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 47 | handle THM _ => Seq.empty; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 48 | |
| 30541 | 49 | fun case_tac ctxt s = gen_case_tac ctxt s NONE; | 
| 50 | fun case_rule_tac ctxt s rule = gen_case_tac ctxt s (SOME rule); | |
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 51 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 52 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 53 | (* induction *) | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 54 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 55 | local | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 56 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 57 | fun prep_var (Var (ixn, _), SOME x) = SOME (ixn, x) | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 58 | | prep_var _ = NONE; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 59 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 60 | fun prep_inst (concl, xs) = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 61 | let val vs = Induct.vars_of concl | 
| 33957 | 62 | in map_filter prep_var (drop (length vs - length xs) vs ~~ xs) end; | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 63 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 64 | in | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 65 | |
| 30541 | 66 | fun gen_induct_tac ctxt0 varss opt_rules i st = | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 67 | let | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 68 | val ((_, goal), ctxt) = Variable.focus_subgoal i st ctxt0; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 69 | val (prems, concl) = Logic.strip_horn (Thm.term_of goal); | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 70 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 71 | fun induct_var name = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 72 | let | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 73 | val t = Syntax.read_term ctxt name; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 74 | val (x, _) = Term.dest_Free t handle TERM _ => | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 75 |           error ("Induction argument not a variable: " ^ Syntax.string_of_term ctxt t);
 | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 76 | val eq_x = fn Free (y, _) => x = y | _ => false; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 77 | val _ = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 78 | if Term.exists_subterm eq_x concl then () | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 79 |           else error ("No such variable in subgoal: " ^ Syntax.string_of_term ctxt t);
 | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 80 | val _ = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 81 | if (exists o Term.exists_subterm) eq_x prems then | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 82 |             warning ("Induction variable occurs also among premises: " ^ Syntax.string_of_term ctxt t)
 | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 83 | else (); | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 84 | in #1 (check_type ctxt t) end; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 85 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 86 | val argss = map (map (Option.map induct_var)) varss; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 87 | val rule = | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 88 | (case opt_rules of | 
| 33368 | 89 | SOME rules => #2 (Rule_Cases.strict_mutual_rule ctxt rules) | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 90 | | NONE => | 
| 33368 | 91 | (case map_filter (Rule_Cases.mutual_rule ctxt) (Induct.get_inductT ctxt argss) of | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 92 | (_, rule) :: _ => rule | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 93 |           | [] => raise THM ("No induction rules", 0, [])));
 | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 94 | |
| 35625 | 95 | val rule' = rule |> Conv.fconv_rule (Conv.concl_conv ~1 Object_Logic.atomize); | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 96 | val _ = Method.trace ctxt [rule']; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 97 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 98 | val concls = Logic.dest_conjunctions (Thm.concl_of rule); | 
| 40722 
441260986b63
make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
 wenzelm parents: 
36960diff
changeset | 99 | val insts = maps prep_inst (concls ~~ varss) handle ListPair.UnequalLengths => | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 100 | error "Induction rule has different numbers of variables"; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 101 | in res_inst_tac ctxt insts rule' i st end | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 102 | handle THM _ => Seq.empty; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 103 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 104 | end; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 105 | |
| 30541 | 106 | fun induct_tac ctxt args = gen_induct_tac ctxt args NONE; | 
| 107 | fun induct_rules_tac ctxt args rules = gen_induct_tac ctxt args (SOME rules); | |
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 108 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 109 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 110 | (* method syntax *) | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 111 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 112 | local | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 113 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 114 | val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.$$$ ":"); | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 115 | val opt_rule = Scan.option (rule_spec |-- Attrib.thm); | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 116 | val opt_rules = Scan.option (rule_spec |-- Attrib.thms); | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 117 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 118 | val varss = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
35625diff
changeset | 119 | Parse.and_list' (Scan.repeat (Scan.unless rule_spec (Scan.lift (Args.maybe Args.name_source)))); | 
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 120 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 121 | in | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 122 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 123 | val setup = | 
| 30541 | 124 |   Method.setup @{binding case_tac}
 | 
| 125 | (Args.goal_spec -- Scan.lift Args.name_source -- opt_rule >> | |
| 126 | (fn ((quant, s), r) => fn ctxt => SIMPLE_METHOD'' quant (gen_case_tac ctxt s r))) | |
| 127 | "unstructured case analysis on types" #> | |
| 128 |   Method.setup @{binding induct_tac}
 | |
| 129 | (Args.goal_spec -- varss -- opt_rules >> | |
| 130 | (fn ((quant, vs), rs) => fn ctxt => SIMPLE_METHOD'' quant (gen_induct_tac ctxt vs rs))) | |
| 131 | "unstructured induction on types"; | |
| 27326 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 132 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 133 | end; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 134 | |
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 135 | end; | 
| 
d3beec370964
moved src/HOL/Tools/induct_tacs.ML to src/Tools/induct_tacs.ML;
 wenzelm parents: diff
changeset | 136 |