1 (* Title: Tools/induct.ML
2 Author: Markus Wenzel, TU Muenchen
4 Proof by cases, induction, and coinduction.
7 signature INDUCT_DATA =
12 val rulify_fallback: thm list
18 val vars_of: term -> term list
19 val dest_rules: Proof.context ->
20 {type_cases: (string * thm) list, pred_cases: (string * thm) list,
21 type_induct: (string * thm) list, pred_induct: (string * thm) list,
22 type_coinduct: (string * thm) list, pred_coinduct: (string * thm) list}
23 val print_rules: Proof.context -> unit
24 val lookup_casesT: Proof.context -> string -> thm option
25 val lookup_casesP: Proof.context -> string -> thm option
26 val lookup_inductT: Proof.context -> string -> thm option
27 val lookup_inductP: Proof.context -> string -> thm option
28 val lookup_coinductT: Proof.context -> string -> thm option
29 val lookup_coinductP: Proof.context -> string -> thm option
30 val find_casesT: Proof.context -> typ -> thm list
31 val find_casesP: Proof.context -> term -> thm list
32 val find_inductT: Proof.context -> typ -> thm list
33 val find_inductP: Proof.context -> term -> thm list
34 val find_coinductT: Proof.context -> typ -> thm list
35 val find_coinductP: Proof.context -> term -> thm list
36 val cases_type: string -> attribute
37 val cases_pred: string -> attribute
38 val cases_del: attribute
39 val induct_type: string -> attribute
40 val induct_pred: string -> attribute
41 val induct_del: attribute
42 val coinduct_type: string -> attribute
43 val coinduct_pred: string -> attribute
44 val coinduct_del: attribute
52 val fix_tac: Proof.context -> int -> (string * typ) list -> int -> tactic
53 val add_defs: (binding option * term) option list -> Proof.context ->
54 (term option list * thm list) * Proof.context
55 val atomize_term: theory -> term -> term
56 val atomize_tac: int -> tactic
57 val inner_atomize_tac: int -> tactic
58 val rulified_term: thm -> theory * term
59 val rulify_tac: int -> tactic
60 val internalize: int -> thm -> thm
61 val guess_instance: Proof.context -> thm -> int -> thm -> thm Seq.seq
62 val cases_tac: Proof.context -> term option list list -> thm option ->
63 thm list -> int -> cases_tactic
64 val get_inductT: Proof.context -> term option list list -> thm list list
65 val induct_tac: Proof.context -> (binding option * term) option list list ->
66 (string * typ) list list -> term option list -> thm list option ->
67 thm list -> int -> cases_tactic
68 val coinduct_tac: Proof.context -> term option list -> term option list -> thm option ->
69 thm list -> int -> cases_tactic
70 val setup: theory -> theory
73 functor Induct(Data: INDUCT_DATA): INDUCT =
79 (* encode_type -- for indexing purposes *)
81 fun encode_type (Type (c, Ts)) = Term.list_comb (Const (c, dummyT), map encode_type Ts)
82 | encode_type (TFree (a, _)) = Free (a, dummyT)
83 | encode_type (TVar (a, _)) = Var (a, dummyT);
86 (* variables -- ordered left-to-right, preferring right *)
89 rev (distinct (op =) (Term.fold_aterms (fn (t as Var _) => cons t | _ => I) tm []));
93 val mk_var = encode_type o #2 o Term.dest_Var;
95 fun concl_var which thm = mk_var (which (vars_of (Thm.concl_of thm))) handle Empty =>
96 raise THM ("No variables in conclusion of rule", 0, [thm]);
100 fun left_var_prem thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle Empty =>
101 raise THM ("No variables in major premise of rule", 0, [thm]);
103 val left_var_concl = concl_var hd;
104 val right_var_concl = concl_var List.last;
114 type rules = (string * thm) Item_Net.T;
117 Item_Net.init (fn ((s1: string, th1), (s2, th2)) => s1 = s2 andalso
118 Thm.eq_thm_prop (th1, th2));
120 fun filter_rules (rs: rules) th =
121 filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs);
123 fun lookup_rule (rs: rules) = AList.lookup (op =) (Item_Net.content rs);
125 fun pretty_rules ctxt kind rs =
126 let val thms = map snd (Item_Net.content rs)
127 in Pretty.big_list kind (map (Display.pretty_thm ctxt) thms) end;
132 structure InductData = GenericDataFun
134 type T = (rules * rules) * (rules * rules) * (rules * rules);
136 ((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)),
137 (init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)),
138 (init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2)));
140 fun merge _ (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1)),
141 ((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2))) =
142 ((Item_Net.merge (casesT1, casesT2), Item_Net.merge (casesP1, casesP2)),
143 (Item_Net.merge (inductT1, inductT2), Item_Net.merge (inductP1, inductP2)),
144 (Item_Net.merge (coinductT1, coinductT2), Item_Net.merge (coinductP1, coinductP2)));
147 val get_local = InductData.get o Context.Proof;
149 fun dest_rules ctxt =
150 let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP)) = get_local ctxt in
151 {type_cases = Item_Net.content casesT,
152 pred_cases = Item_Net.content casesP,
153 type_induct = Item_Net.content inductT,
154 pred_induct = Item_Net.content inductP,
155 type_coinduct = Item_Net.content coinductT,
156 pred_coinduct = Item_Net.content coinductP}
159 fun print_rules ctxt =
160 let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP)) = get_local ctxt in
161 [pretty_rules ctxt "coinduct type:" coinductT,
162 pretty_rules ctxt "coinduct pred:" coinductP,
163 pretty_rules ctxt "induct type:" inductT,
164 pretty_rules ctxt "induct pred:" inductP,
165 pretty_rules ctxt "cases type:" casesT,
166 pretty_rules ctxt "cases pred:" casesP]
167 |> Pretty.chunks |> Pretty.writeln
171 OuterSyntax.improper_command "print_induct_rules" "print induction and cases rules"
172 OuterKeyword.diag (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
173 Toplevel.keep (print_rules o Toplevel.context_of)));
178 val lookup_casesT = lookup_rule o #1 o #1 o get_local;
179 val lookup_casesP = lookup_rule o #2 o #1 o get_local;
180 val lookup_inductT = lookup_rule o #1 o #2 o get_local;
181 val lookup_inductP = lookup_rule o #2 o #2 o get_local;
182 val lookup_coinductT = lookup_rule o #1 o #3 o get_local;
183 val lookup_coinductP = lookup_rule o #2 o #3 o get_local;
186 fun find_rules which how ctxt x =
187 map snd (Item_Net.retrieve (which (get_local ctxt)) (how x));
189 val find_casesT = find_rules (#1 o #1) encode_type;
190 val find_casesP = find_rules (#2 o #1) I;
191 val find_inductT = find_rules (#1 o #2) encode_type;
192 val find_inductP = find_rules (#2 o #2) I;
193 val find_coinductT = find_rules (#1 o #3) encode_type;
194 val find_coinductP = find_rules (#2 o #3) I;
202 fun mk_att f g name arg =
203 let val (x, thm) = g arg in (InductData.map (f (name, thm)) x, thm) end;
205 fun del_att which = Thm.declaration_attribute (fn th => InductData.map (which (pairself (fn rs =>
206 fold Item_Net.delete (filter_rules rs th) rs))));
208 fun map1 f (x, y, z) = (f x, y, z);
209 fun map2 f (x, y, z) = (x, f y, z);
210 fun map3 f (x, y, z) = (x, y, f z);
212 fun add_casesT rule x = map1 (apfst (Item_Net.insert rule)) x;
213 fun add_casesP rule x = map1 (apsnd (Item_Net.insert rule)) x;
214 fun add_inductT rule x = map2 (apfst (Item_Net.insert rule)) x;
215 fun add_inductP rule x = map2 (apsnd (Item_Net.insert rule)) x;
216 fun add_coinductT rule x = map3 (apfst (Item_Net.insert rule)) x;
217 fun add_coinductP rule x = map3 (apsnd (Item_Net.insert rule)) x;
219 val consumes0 = Rule_Cases.consumes_default 0;
220 val consumes1 = Rule_Cases.consumes_default 1;
224 val cases_type = mk_att add_casesT consumes0;
225 val cases_pred = mk_att add_casesP consumes1;
226 val cases_del = del_att map1;
228 val induct_type = mk_att add_inductT consumes0;
229 val induct_pred = mk_att add_inductP consumes1;
230 val induct_del = del_att map2;
232 val coinduct_type = mk_att add_coinductT consumes0;
233 val coinduct_pred = mk_att add_coinductP consumes1;
234 val coinduct_del = del_att map3;
240 (** attribute syntax **)
242 val casesN = "cases";
243 val inductN = "induct";
244 val coinductN = "coinduct";
253 Scan.lift (Args.$$$ k -- Args.colon) |-- arg ||
254 Scan.lift (Args.$$$ k) >> K "";
256 fun attrib add_type add_pred del =
257 spec typeN Args.tyname >> add_type ||
258 spec predN Args.const >> add_pred ||
259 spec setN Args.const >> add_pred ||
260 Scan.lift Args.del >> K del;
265 Attrib.setup @{binding cases} (attrib cases_type cases_pred cases_del)
266 "declaration of cases rule" #>
267 Attrib.setup @{binding induct} (attrib induct_type induct_pred induct_del)
268 "declaration of induction rule" #>
269 Attrib.setup @{binding coinduct} (attrib coinduct_type coinduct_pred coinduct_del)
270 "declaration of coinduction rule";
280 fun align_left msg xs ys =
281 let val m = length xs and n = length ys
282 in if m < n then error msg else (Library.take (n, xs) ~~ ys) end;
284 fun align_right msg xs ys =
285 let val m = length xs and n = length ys
286 in if m < n then error msg else (Library.drop (m - n, xs) ~~ ys) end;
291 fun prep_inst ctxt align tune (tm, ts) =
293 val cert = Thm.cterm_of (ProofContext.theory_of ctxt);
294 fun prep_var (x, SOME t) =
297 val xT = #T (Thm.rep_cterm cx);
298 val ct = cert (tune t);
299 val tT = #T (Thm.rep_cterm ct);
301 if Type.could_unify (tT, xT) then SOME (cx, ct)
302 else error (Pretty.string_of (Pretty.block
303 [Pretty.str "Ill-typed instantiation:", Pretty.fbrk,
304 Syntax.pretty_term ctxt (Thm.term_of ct), Pretty.str " ::", Pretty.brk 1,
305 Syntax.pretty_typ ctxt tT]))
307 | prep_var (_, NONE) = NONE;
310 align "Rule has fewer variables than instantiations given" xs ts
311 |> map_filter prep_var
317 fun trace_rules _ kind [] = error ("Unable to figure out " ^ kind ^ " rule")
318 | trace_rules ctxt _ rules = Method.trace ctxt rules;
325 rule selection scheme:
326 cases - default case split
327 `A t` cases ... - predicate/set cases
329 ... cases ... r - explicit rule
334 fun get_casesT ctxt ((SOME t :: _) :: _) = find_casesT ctxt (Term.fastype_of t)
335 | get_casesT _ _ = [];
337 fun get_casesP ctxt (fact :: _) = find_casesP ctxt (Thm.concl_of fact)
338 | get_casesP _ _ = [];
342 fun cases_tac ctxt insts opt_rule facts =
344 val thy = ProofContext.theory_of ctxt;
347 if null insts then `Rule_Cases.get r
348 else (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts
349 |> maps (prep_inst ctxt align_left I)
350 |> Drule.cterm_instantiate) r |> pair (Rule_Cases.get r);
354 SOME r => Seq.single (inst_rule r)
356 (get_casesP ctxt facts @ get_casesT ctxt insts @ [Data.cases_default])
357 |> tap (trace_rules ctxt casesN)
358 |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
362 |> Seq.maps (Rule_Cases.consume [] facts)
363 |> Seq.maps (fn ((cases, (_, more_facts)), rule) =>
364 CASES (Rule_Cases.make_common false (thy, Thm.prop_of rule) cases)
365 (Method.insert_tac more_facts i THEN Tactic.rtac rule i) st)
372 (** induct method **)
374 val conjunction_congs = [@{thm Pure.all_conjunction}, @{thm imp_conjunction}];
379 fun atomize_term thy =
380 MetaSimplifier.rewrite_term thy Data.atomize []
381 #> ObjectLogic.drop_judgment thy;
383 val atomize_cterm = MetaSimplifier.rewrite true Data.atomize;
385 val atomize_tac = Simplifier.rewrite_goal_tac Data.atomize;
387 val inner_atomize_tac =
388 Simplifier.rewrite_goal_tac (map Thm.symmetric conjunction_congs) THEN' atomize_tac;
393 fun rulify_term thy =
394 MetaSimplifier.rewrite_term thy (Data.rulify @ conjunction_congs) [] #>
395 MetaSimplifier.rewrite_term thy Data.rulify_fallback [];
397 fun rulified_term thm =
399 val thy = Thm.theory_of_thm thm;
400 val rulify = rulify_term thy;
401 val (As, B) = Logic.strip_horn (Thm.prop_of thm);
402 in (thy, Logic.list_implies (map rulify As, rulify B)) end;
405 Simplifier.rewrite_goal_tac (Data.rulify @ conjunction_congs) THEN'
406 Simplifier.rewrite_goal_tac Data.rulify_fallback THEN'
407 Goal.conjunction_tac THEN_ALL_NEW
408 (Simplifier.rewrite_goal_tac [@{thm Pure.conjunction_imp}] THEN' Goal.norm_hhf_tac);
413 fun rule_instance ctxt inst rule =
414 Drule.cterm_instantiate (prep_inst ctxt align_left I (Thm.prop_of rule, inst)) rule;
416 fun internalize k th =
417 th |> Thm.permute_prems 0 k
418 |> Conv.fconv_rule (Conv.concl_conv (Thm.nprems_of th - k) atomize_cterm);
421 (* guess rule instantiation -- cannot handle pending goal parameters *)
425 fun dest_env thy env =
427 val cert = Thm.cterm_of thy;
428 val certT = Thm.ctyp_of thy;
429 val pairs = Vartab.dest (Envir.term_env env);
430 val types = Vartab.dest (Envir.type_env env);
431 val ts = map (cert o Envir.norm_term env o #2 o #2) pairs;
432 val xs = map2 (curry (cert o Var)) (map #1 pairs) (map (#T o Thm.rep_cterm) ts);
433 in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) types, xs ~~ ts) end;
437 fun guess_instance ctxt rule i st =
439 val thy = ProofContext.theory_of ctxt;
440 val maxidx = Thm.maxidx_of st;
441 val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*)
442 val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal));
444 if not (null params) then
445 (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^
446 commas_quote (map (Syntax.string_of_term ctxt o Syntax.mark_boundT) params));
450 val rule' = Thm.incr_indexes (maxidx + 1) rule;
451 val concl = Logic.strip_assums_concl goal;
453 Unify.smash_unifiers thy [(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule'))
454 |> Seq.map (fn env => Drule.instantiate (dest_env thy env) rule')
456 end handle Subscript => Seq.empty;
461 (* special renaming of rule parameters *)
463 fun special_rename_params ctxt [[SOME (Free (z, Type (T, _)))]] [thm] =
465 val x = Name.clean (ProofContext.revert_skolem ctxt z);
467 | index i (y :: ys) =
468 if x = y then x ^ string_of_int i :: index (i + 1) ys
469 else y :: index i ys;
470 fun rename_params [] = []
471 | rename_params ((y, Type (U, _)) :: ys) =
472 (if U = T then x else y) :: rename_params ys
473 | rename_params ((y, _) :: ys) = y :: rename_params ys;
476 val xs = rename_params (Logic.strip_params A);
478 (case filter (fn x' => x' = x) xs of
479 [] => xs | [_] => xs | _ => index 1 xs);
480 in Logic.list_rename_params (xs', A) end;
482 let val (As, C) = Logic.strip_horn p
483 in Logic.list_implies (map rename_asm As, C) end;
484 val cp' = cterm_fun rename_prop (Thm.cprop_of thm);
485 val thm' = Thm.equal_elim (Thm.reflexive cp') thm;
486 in [Rule_Cases.save thm thm'] end
487 | special_rename_params _ _ ths = ths;
494 fun goal_prefix k ((c as Const ("all", _)) $ Abs (a, T, B)) = c $ Abs (a, T, goal_prefix k B)
495 | goal_prefix 0 _ = Term.dummy_pattern propT
496 | goal_prefix k ((c as Const ("==>", _)) $ A $ B) = c $ A $ goal_prefix (k - 1) B
497 | goal_prefix _ _ = Term.dummy_pattern propT;
499 fun goal_params k (Const ("all", _) $ Abs (_, _, B)) = goal_params k B + 1
500 | goal_params 0 _ = 0
501 | goal_params k (Const ("==>", _) $ _ $ B) = goal_params (k - 1) B
502 | goal_params _ _ = 0;
504 fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) =>
506 val thy = ProofContext.theory_of ctxt;
507 val cert = Thm.cterm_of thy;
510 fun spec_rule prfx (xs, body) =
511 @{thm Pure.meta_spec}
512 |> Thm.rename_params_rule ([Name.clean (ProofContext.revert_skolem ctxt x)], 1)
513 |> Thm.lift_rule (cert prfx)
514 |> `(Thm.prop_of #> Logic.strip_assums_concl)
515 |-> (fn pred $ arg =>
516 Drule.cterm_instantiate
517 [(cert (Term.head_of pred), cert (Logic.rlist_abs (xs, body))),
518 (cert (Term.head_of arg), cert (Logic.rlist_abs (xs, v)))]);
520 fun goal_concl k xs (Const ("all", _) $ Abs (a, T, B)) = goal_concl k ((a, T) :: xs) B
521 | goal_concl 0 xs B =
522 if not (Term.exists_subterm (fn t => t aconv v) B) then NONE
523 else SOME (xs, Term.absfree (x, T, Term.incr_boundvars 1 B))
524 | goal_concl k xs (Const ("==>", _) $ _ $ B) = goal_concl (k - 1) xs B
525 | goal_concl _ _ _ = NONE;
527 (case goal_concl n [] goal of
529 (compose_tac (false, spec_rule (goal_prefix n goal) concl, 1) THEN' rtac asm_rl) i
533 fun miniscope_tac p = CONVERSION o
534 Conv.params_conv p (K (MetaSimplifier.rewrite true [Thm.symmetric Drule.norm_hhf_eq]));
538 fun fix_tac _ _ [] = K all_tac
539 | fix_tac ctxt n xs = SUBGOAL (fn (goal, i) =>
540 (EVERY' (map (meta_spec_tac ctxt n) xs) THEN'
541 (miniscope_tac (goal_params n goal) ctxt)) i);
548 fun add_defs def_insts =
550 fun add (SOME (SOME x, t)) ctxt =
551 let val ([(lhs, (_, th))], ctxt') =
552 LocalDefs.add_defs [((x, NoSyn), (Thm.empty_binding, t))] ctxt
553 in ((SOME lhs, [th]), ctxt') end
554 | add (SOME (NONE, t)) ctxt = ((SOME t, []), ctxt)
555 | add NONE ctxt = ((NONE, []), ctxt);
556 in fold_map add def_insts #> apfst (split_list #> apsnd flat) end;
562 rule selection scheme:
563 `A x` induct ... - predicate/set induction
564 induct x - type induction
565 ... induct ... r - explicit rule
568 fun get_inductT ctxt insts =
569 fold_rev (map_product cons) (insts |> map
570 ((fn [] => NONE | ts => List.last ts) #>
571 (fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #>
572 find_inductT ctxt)) [[]]
573 |> filter_out (forall Rule_Cases.is_inner_rule);
575 fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact))
576 | get_inductP _ _ = [];
578 fun induct_tac ctxt def_insts arbitrary taking opt_rule facts =
580 val thy = ProofContext.theory_of ctxt;
582 val ((insts, defs), defs_ctxt) = fold_map add_defs def_insts ctxt |>> split_list;
583 val atomized_defs = map (map (Conv.fconv_rule ObjectLogic.atomize)) defs;
585 fun inst_rule (concls, r) =
586 (if null insts then `Rule_Cases.get r
587 else (align_left "Rule has fewer conclusions than arguments given"
588 (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts
589 |> maps (prep_inst ctxt align_right (atomize_term thy))
590 |> Drule.cterm_instantiate) r |> pair (Rule_Cases.get r))
591 |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th));
595 SOME rs => Seq.single (inst_rule (Rule_Cases.strict_mutual_rule ctxt rs))
597 (get_inductP ctxt facts @
598 map (special_rename_params defs_ctxt insts) (get_inductT ctxt insts))
599 |> map_filter (Rule_Cases.mutual_rule ctxt)
600 |> tap (trace_rules ctxt inductN o map #2)
601 |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
603 fun rule_cases rule =
604 Rule_Cases.make_nested false (Thm.prop_of rule) (rulified_term rule);
608 |> Seq.maps (Rule_Cases.consume (flat defs) facts)
609 |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) =>
610 (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j =>
612 (Method.insert_tac (more_facts @ nth_list atomized_defs (j - 1))
613 THEN' fix_tac defs_ctxt
614 (nth concls (j - 1) + more_consumes)
615 (nth_list arbitrary (j - 1))))
616 THEN' inner_atomize_tac) j))
617 THEN' atomize_tac) i st |> Seq.maps (fn st' =>
618 guess_instance ctxt (internalize more_consumes rule) i st'
619 |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
620 |> Seq.maps (fn rule' =>
621 CASES (rule_cases rule' cases)
622 (Tactic.rtac rule' i THEN
623 PRIMITIVE (singleton (ProofContext.export defs_ctxt ctxt))) st'))))
624 THEN_ALL_NEW_CASES rulify_tac
629 (** coinduct method **)
632 rule selection scheme:
633 goal "A x" coinduct ... - predicate/set coinduction
634 coinduct x - type coinduction
635 coinduct ... r - explicit rule
640 fun get_coinductT ctxt (SOME t :: _) = find_coinductT ctxt (Term.fastype_of t)
641 | get_coinductT _ _ = [];
643 fun get_coinductP ctxt goal = find_coinductP ctxt (Logic.strip_assums_concl goal);
645 fun main_prop_of th =
646 if Rule_Cases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th;
650 fun coinduct_tac ctxt inst taking opt_rule facts =
652 val thy = ProofContext.theory_of ctxt;
655 if null inst then `Rule_Cases.get r
656 else Drule.cterm_instantiate (prep_inst ctxt align_right I (main_prop_of r, inst)) r
657 |> pair (Rule_Cases.get r);
661 SOME r => Seq.single (inst_rule r)
663 (get_coinductP ctxt goal @ get_coinductT ctxt inst)
664 |> tap (trace_rules ctxt coinductN)
665 |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
667 SUBGOAL_CASES (fn (goal, i) => fn st =>
669 |> Seq.maps (Rule_Cases.consume [] facts)
670 |> Seq.maps (fn ((cases, (_, more_facts)), rule) =>
671 guess_instance ctxt rule i st
672 |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
673 |> Seq.maps (fn rule' =>
674 CASES (Rule_Cases.make_common false (thy, Thm.prop_of rule') cases)
675 (Method.insert_tac more_facts i THEN Tactic.rtac rule' i) st)))
682 (** concrete syntax **)
684 structure P = OuterParse;
686 val arbitraryN = "arbitrary";
687 val takingN = "taking";
692 fun single_rule [rule] = rule
693 | single_rule _ = error "Single rule expected";
695 fun named_rule k arg get =
696 Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|--
697 (fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name =>
698 (case get (Context.proof_of context) name of SOME x => x
699 | NONE => error ("No rule for " ^ k ^ " " ^ quote name))))));
701 fun rule get_type get_pred =
702 named_rule typeN Args.tyname get_type ||
703 named_rule predN Args.const get_pred ||
704 named_rule setN Args.const get_pred ||
705 Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms;
707 val cases_rule = rule lookup_casesT lookup_casesP >> single_rule;
708 val induct_rule = rule lookup_inductT lookup_inductP;
709 val coinduct_rule = rule lookup_coinductT lookup_coinductP >> single_rule;
711 val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME;
714 ((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME)
715 -- Args.term) >> SOME ||
716 inst >> Option.map (pair NONE);
718 val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) =>
719 error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
721 fun unless_more_args scan = Scan.unless (Scan.lift
722 ((Args.$$$ arbitraryN || Args.$$$ takingN || Args.$$$ typeN ||
723 Args.$$$ predN || Args.$$$ setN || Args.$$$ ruleN) -- Args.colon)) scan;
725 val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |--
726 P.and_list1' (Scan.repeat (unless_more_args free))) [];
728 val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |--
729 Scan.repeat1 (unless_more_args inst)) [];
734 Method.setup @{binding cases}
735 (P.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule >>
736 (fn (insts, opt_rule) => fn ctxt =>
737 METHOD_CASES (fn facts => Seq.DETERM (HEADGOAL (cases_tac ctxt insts opt_rule facts)))))
738 "case analysis on types or predicates/sets";
741 Method.setup @{binding induct}
742 (P.and_list' (Scan.repeat (unless_more_args def_inst)) --
743 (arbitrary -- taking -- Scan.option induct_rule) >>
744 (fn (insts, ((arbitrary, taking), opt_rule)) => fn ctxt =>
745 RAW_METHOD_CASES (fn facts =>
746 Seq.DETERM (HEADGOAL (induct_tac ctxt insts arbitrary taking opt_rule facts)))))
747 "induction on types or predicates/sets";
750 Method.setup @{binding coinduct}
751 (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >>
752 (fn ((insts, taking), opt_rule) => fn ctxt =>
753 RAW_METHOD_CASES (fn facts =>
754 Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts)))))
755 "coinduction on types or predicates/sets";
763 val setup = attrib_setup #> cases_setup #> induct_setup #> coinduct_setup;