1 (* Title: Tools/induct.ML
2 Author: Markus Wenzel, TU Muenchen
4 Proof by cases, induction, and coinduction.
7 signature INDUCT_ARGS =
12 val rulify_fallback: thm list
14 val dest_def: term -> (term * term) option
15 val trivial_tac: int -> tactic
21 val vars_of: term -> term list
22 val dest_rules: Proof.context ->
23 {type_cases: (string * thm) list, pred_cases: (string * thm) list,
24 type_induct: (string * thm) list, pred_induct: (string * thm) list,
25 type_coinduct: (string * thm) list, pred_coinduct: (string * thm) list}
26 val print_rules: Proof.context -> unit
27 val lookup_casesT: Proof.context -> string -> thm option
28 val lookup_casesP: Proof.context -> string -> thm option
29 val lookup_inductT: Proof.context -> string -> thm option
30 val lookup_inductP: Proof.context -> string -> thm option
31 val lookup_coinductT: Proof.context -> string -> thm option
32 val lookup_coinductP: Proof.context -> string -> thm option
33 val find_casesT: Proof.context -> typ -> thm list
34 val find_casesP: Proof.context -> term -> thm list
35 val find_inductT: Proof.context -> typ -> thm list
36 val find_inductP: Proof.context -> term -> thm list
37 val find_coinductT: Proof.context -> typ -> thm list
38 val find_coinductP: Proof.context -> term -> thm list
39 val cases_type: string -> attribute
40 val cases_pred: string -> attribute
41 val cases_del: attribute
42 val induct_type: string -> attribute
43 val induct_pred: string -> attribute
44 val induct_del: attribute
45 val coinduct_type: string -> attribute
46 val coinduct_pred: string -> attribute
47 val coinduct_del: attribute
48 val map_simpset: (simpset -> simpset) -> Context.generic -> Context.generic
49 val induct_simp_add: attribute
50 val induct_simp_del: attribute
59 val arbitrary_tac: Proof.context -> int -> (string * typ) list -> int -> tactic
60 val add_defs: (binding option * (term * bool)) option list -> Proof.context ->
61 (term option list * thm list) * Proof.context
62 val atomize_term: theory -> term -> term
63 val atomize_cterm: conv
64 val atomize_tac: int -> tactic
65 val inner_atomize_tac: int -> tactic
66 val rulified_term: thm -> theory * term
67 val rulify_tac: int -> tactic
68 val simplified_rule: Proof.context -> thm -> thm
69 val simplify_tac: Proof.context -> int -> tactic
70 val trivial_tac: int -> tactic
71 val rotate_tac: int -> int -> int -> tactic
72 val internalize: int -> thm -> thm
73 val guess_instance: Proof.context -> thm -> int -> thm -> thm Seq.seq
74 val cases_tac: Proof.context -> bool -> term option list list -> thm option ->
75 thm list -> int -> cases_tactic
76 val get_inductT: Proof.context -> term option list list -> thm list list
77 type case_data = (((string * string list) * string list) list * int) (* FIXME -> rule_cases.ML *)
78 val gen_induct_tac: (theory -> case_data * thm -> case_data * thm) ->
79 Proof.context -> bool -> (binding option * (term * bool)) option list list ->
80 (string * typ) list list -> term option list -> thm list option ->
81 thm list -> int -> cases_tactic
82 val induct_tac: Proof.context -> bool -> (binding option * (term * bool)) option list list ->
83 (string * typ) list list -> term option list -> thm list option ->
84 thm list -> int -> cases_tactic
85 val coinduct_tac: Proof.context -> term option list -> term option list -> thm option ->
86 thm list -> int -> cases_tactic
87 val gen_induct_setup: binding ->
88 (Proof.context -> bool -> (binding option * (term * bool)) option list list ->
89 (string * typ) list list -> term option list -> thm list option ->
90 thm list -> int -> cases_tactic) ->
92 val setup: theory -> theory
95 functor Induct(Induct_Args: INDUCT_ARGS): INDUCT =
98 (** variables -- ordered left-to-right, preferring right **)
101 rev (distinct (op =) (Term.fold_aterms (fn t as Var _ => cons t | _ => I) tm []));
105 val mk_var = Net.encode_type o #2 o Term.dest_Var;
107 fun concl_var which thm = mk_var (which (vars_of (Thm.concl_of thm))) handle List.Empty =>
108 raise THM ("No variables in conclusion of rule", 0, [thm]);
112 fun left_var_prem thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle List.Empty =>
113 raise THM ("No variables in major premise of rule", 0, [thm]);
115 val left_var_concl = concl_var hd;
116 val right_var_concl = concl_var List.last;
122 (** constraint simplification **)
124 (* rearrange parameters and premises to allow application of one-point-rules *)
126 fun swap_params_conv ctxt i j cv =
128 fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt
130 Conv.rewr_conv @{thm swap_params} then_conv
131 Conv.forall_conv (conv1 (k - 1) o snd) ctxt
132 fun conv2 0 ctxt = conv1 j ctxt
133 | conv2 k ctxt = Conv.forall_conv (conv2 (k - 1) o snd) ctxt
136 fun swap_prems_conv 0 = Conv.all_conv
137 | swap_prems_conv i =
138 Conv.implies_concl_conv (swap_prems_conv (i - 1)) then_conv
139 Conv.rewr_conv Drule.swap_prems_eq
141 fun drop_judgment ctxt = Object_Logic.drop_judgment (Proof_Context.theory_of ctxt);
145 val l = length (Logic.strip_params t);
146 val Hs = Logic.strip_assums_hyp t;
148 (case Induct_Args.dest_def (drop_judgment ctxt t) of
149 SOME (Bound j, _) => SOME (i, j)
150 | SOME (_, Bound j) => SOME (i, j)
153 (case get_first find (map_index I Hs) of
155 | SOME (0, 0) => NONE
156 | SOME (i, j) => SOME (i, l - j - 1, j))
159 fun mk_swap_rrule ctxt ct =
160 (case find_eq ctxt (term_of ct) of
162 | SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct));
164 val rearrange_eqs_simproc =
165 Simplifier.simproc_global
166 (Thm.theory_of_thm Drule.swap_prems_eq) "rearrange_eqs" ["all t"]
167 (fn thy => fn ss => fn t =>
168 mk_swap_rrule (Simplifier.the_context ss) (cterm_of thy t));
171 (* rotate k premises to the left by j, skipping over first j premises *)
173 fun rotate_conv 0 j 0 = Conv.all_conv
174 | rotate_conv 0 j k = swap_prems_conv j then_conv rotate_conv 1 j (k - 1)
175 | rotate_conv i j k = Conv.implies_concl_conv (rotate_conv (i - 1) j k);
177 fun rotate_tac j 0 = K all_tac
178 | rotate_tac j k = SUBGOAL (fn (goal, i) =>
179 CONVERSION (rotate_conv
180 j (length (Logic.strip_assums_hyp goal) - j - k) k) i);
183 (* rulify operators around definition *)
185 fun rulify_defs_conv ctxt ct =
186 if exists_subterm (is_some o Induct_Args.dest_def) (term_of ct) andalso
187 not (is_some (Induct_Args.dest_def (drop_judgment ctxt (term_of ct))))
189 (Conv.forall_conv (rulify_defs_conv o snd) ctxt else_conv
190 Conv.implies_conv (Conv.try_conv (rulify_defs_conv ctxt))
191 (Conv.try_conv (rulify_defs_conv ctxt)) else_conv
192 Conv.first_conv (map Conv.rewr_conv Induct_Args.rulify) then_conv
193 Conv.try_conv (rulify_defs_conv ctxt)) ct
194 else Conv.no_conv ct;
202 type rules = (string * thm) Item_Net.T;
204 fun init_rules index : rules =
206 (fn ((s1, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm_prop (th1, th2))
209 fun filter_rules (rs: rules) th =
210 filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs);
212 fun lookup_rule (rs: rules) = AList.lookup (op =) (Item_Net.content rs);
214 fun pretty_rules ctxt kind rs =
215 let val thms = map snd (Item_Net.content rs)
216 in Pretty.big_list kind (map (Display.pretty_thm ctxt) thms) end;
221 structure Data = Generic_Data
223 type T = (rules * rules) * (rules * rules) * (rules * rules) * simpset;
225 ((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)),
226 (init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)),
227 (init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2)),
228 empty_ss addsimprocs [rearrange_eqs_simproc] addsimps [Drule.norm_hhf_eq]);
230 fun merge (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1), simpset1),
231 ((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2), simpset2)) =
232 ((Item_Net.merge (casesT1, casesT2), Item_Net.merge (casesP1, casesP2)),
233 (Item_Net.merge (inductT1, inductT2), Item_Net.merge (inductP1, inductP2)),
234 (Item_Net.merge (coinductT1, coinductT2), Item_Net.merge (coinductP1, coinductP2)),
235 merge_ss (simpset1, simpset2));
238 val get_local = Data.get o Context.Proof;
240 fun dest_rules ctxt =
241 let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in
242 {type_cases = Item_Net.content casesT,
243 pred_cases = Item_Net.content casesP,
244 type_induct = Item_Net.content inductT,
245 pred_induct = Item_Net.content inductP,
246 type_coinduct = Item_Net.content coinductT,
247 pred_coinduct = Item_Net.content coinductP}
250 fun print_rules ctxt =
251 let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in
252 [pretty_rules ctxt "coinduct type:" coinductT,
253 pretty_rules ctxt "coinduct pred:" coinductP,
254 pretty_rules ctxt "induct type:" inductT,
255 pretty_rules ctxt "induct pred:" inductP,
256 pretty_rules ctxt "cases type:" casesT,
257 pretty_rules ctxt "cases pred:" casesP]
258 |> Pretty.chunks |> Pretty.writeln
262 Outer_Syntax.improper_command @{command_spec "print_induct_rules"}
263 "print induction and cases rules"
264 (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
265 Toplevel.keep (print_rules o Toplevel.context_of)));
270 val lookup_casesT = lookup_rule o #1 o #1 o get_local;
271 val lookup_casesP = lookup_rule o #2 o #1 o get_local;
272 val lookup_inductT = lookup_rule o #1 o #2 o get_local;
273 val lookup_inductP = lookup_rule o #2 o #2 o get_local;
274 val lookup_coinductT = lookup_rule o #1 o #3 o get_local;
275 val lookup_coinductP = lookup_rule o #2 o #3 o get_local;
278 fun find_rules which how ctxt x =
279 map snd (Item_Net.retrieve (which (get_local ctxt)) (how x));
281 val find_casesT = find_rules (#1 o #1) Net.encode_type;
282 val find_casesP = find_rules (#2 o #1) I;
283 val find_inductT = find_rules (#1 o #2) Net.encode_type;
284 val find_inductP = find_rules (#2 o #2) I;
285 val find_coinductT = find_rules (#1 o #3) Net.encode_type;
286 val find_coinductP = find_rules (#2 o #3) I;
294 fun mk_att f g name =
295 Thm.mixed_attribute (fn (context, thm) =>
298 val context' = Data.map (f (name, thm')) context;
299 in (context', thm') end);
302 Thm.declaration_attribute (fn th => Data.map (which (pairself (fn rs =>
303 fold Item_Net.remove (filter_rules rs th) rs))));
305 fun map1 f (x, y, z, s) = (f x, y, z, s);
306 fun map2 f (x, y, z, s) = (x, f y, z, s);
307 fun map3 f (x, y, z, s) = (x, y, f z, s);
308 fun map4 f (x, y, z, s) = (x, y, z, f s);
310 fun add_casesT rule x = map1 (apfst (Item_Net.update rule)) x;
311 fun add_casesP rule x = map1 (apsnd (Item_Net.update rule)) x;
312 fun add_inductT rule x = map2 (apfst (Item_Net.update rule)) x;
313 fun add_inductP rule x = map2 (apsnd (Item_Net.update rule)) x;
314 fun add_coinductT rule x = map3 (apfst (Item_Net.update rule)) x;
315 fun add_coinductP rule x = map3 (apsnd (Item_Net.update rule)) x;
317 val consumes0 = Rule_Cases.default_consumes 0;
318 val consumes1 = Rule_Cases.default_consumes 1;
322 val cases_type = mk_att add_casesT consumes0;
323 val cases_pred = mk_att add_casesP consumes1;
324 val cases_del = del_att map1;
326 val induct_type = mk_att add_inductT consumes0;
327 val induct_pred = mk_att add_inductP consumes1;
328 val induct_del = del_att map2;
330 val coinduct_type = mk_att add_coinductT consumes0;
331 val coinduct_pred = mk_att add_coinductP consumes1;
332 val coinduct_del = del_att map3;
334 fun map_simpset f = Data.map (map4 f);
337 Thm.declaration_attribute (fn thm => fn context =>
339 (Simplifier.with_context (Context.proof_of context) (fn ss => f (ss, [thm]))) context);
341 val induct_simp_add = induct_simp (op addsimps);
342 val induct_simp_del = induct_simp (op delsimps);
348 (** attribute syntax **)
350 val no_simpN = "no_simp";
351 val casesN = "cases";
352 val inductN = "induct";
353 val coinductN = "coinduct";
362 Scan.lift (Args.$$$ k -- Args.colon) |-- arg ||
363 Scan.lift (Args.$$$ k) >> K "";
365 fun attrib add_type add_pred del =
366 spec typeN (Args.type_name false) >> add_type ||
367 spec predN (Args.const false) >> add_pred ||
368 spec setN (Args.const false) >> add_pred ||
369 Scan.lift Args.del >> K del;
374 Attrib.setup @{binding cases} (attrib cases_type cases_pred cases_del)
375 "declaration of cases rule" #>
376 Attrib.setup @{binding induct} (attrib induct_type induct_pred induct_del)
377 "declaration of induction rule" #>
378 Attrib.setup @{binding coinduct} (attrib coinduct_type coinduct_pred coinduct_del)
379 "declaration of coinduction rule" #>
380 Attrib.setup @{binding induct_simp} (Attrib.add_del induct_simp_add induct_simp_del)
381 "declaration of rules for simplifying induction or cases rules";
391 fun align_left msg xs ys =
392 let val m = length xs and n = length ys
393 in if m < n then error msg else (take n xs ~~ ys) end;
395 fun align_right msg xs ys =
396 let val m = length xs and n = length ys
397 in if m < n then error msg else (drop (m - n) xs ~~ ys) end;
402 fun prep_inst ctxt align tune (tm, ts) =
404 val cert = Thm.cterm_of (Proof_Context.theory_of ctxt);
405 fun prep_var (x, SOME t) =
408 val xT = #T (Thm.rep_cterm cx);
409 val ct = cert (tune t);
410 val tT = #T (Thm.rep_cterm ct);
412 if Type.could_unify (tT, xT) then SOME (cx, ct)
413 else error (Pretty.string_of (Pretty.block
414 [Pretty.str "Ill-typed instantiation:", Pretty.fbrk,
415 Syntax.pretty_term ctxt (Thm.term_of ct), Pretty.str " ::", Pretty.brk 1,
416 Syntax.pretty_typ ctxt tT]))
418 | prep_var (_, NONE) = NONE;
421 align "Rule has fewer variables than instantiations given" xs ts
422 |> map_filter prep_var
428 fun trace_rules _ kind [] = error ("Unable to figure out " ^ kind ^ " rule")
429 | trace_rules ctxt _ rules = Method.trace ctxt rules;
432 (* mark equality constraints in cases rule *)
434 val equal_def' = Thm.symmetric Induct_Args.equal_def;
436 fun mark_constraints n ctxt = Conv.fconv_rule
437 (Conv.prems_conv ~1 (Conv.params_conv ~1 (K (Conv.prems_conv n
438 (Raw_Simplifier.rewrite false [equal_def']))) ctxt));
440 val unmark_constraints = Conv.fconv_rule
441 (Raw_Simplifier.rewrite true [Induct_Args.equal_def]);
446 fun simplify_conv' ctxt =
447 Simplifier.full_rewrite (Simplifier.context ctxt (#4 (get_local ctxt)));
449 fun simplify_conv ctxt ct =
450 if exists_subterm (is_some o Induct_Args.dest_def) (term_of ct) then
451 (Conv.try_conv (rulify_defs_conv ctxt) then_conv simplify_conv' ctxt) ct
452 else Conv.all_conv ct;
454 fun gen_simplified_rule cv ctxt =
455 Conv.fconv_rule (Conv.prems_conv ~1 (cv ctxt));
457 val simplified_rule' = gen_simplified_rule simplify_conv';
458 val simplified_rule = gen_simplified_rule simplify_conv;
460 fun simplify_tac ctxt = CONVERSION (simplify_conv ctxt);
462 val trivial_tac = Induct_Args.trivial_tac;
469 rule selection scheme:
470 cases - default case split
471 `A t` cases ... - predicate/set cases
473 ... cases ... r - explicit rule
478 fun get_casesT ctxt ((SOME t :: _) :: _) = find_casesT ctxt (Term.fastype_of t)
479 | get_casesT _ _ = [];
481 fun get_casesP ctxt (fact :: _) = find_casesP ctxt (Thm.concl_of fact)
482 | get_casesP _ _ = [];
486 fun cases_tac ctxt simp insts opt_rule facts =
488 val thy = Proof_Context.theory_of ctxt;
491 (if null insts then r
493 (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts
494 |> maps (prep_inst ctxt align_left I)
495 |> Drule.cterm_instantiate) r)
496 |> simp ? mark_constraints (Rule_Cases.get_constraints r) ctxt
497 |> pair (Rule_Cases.get r);
501 SOME r => Seq.single (inst_rule r)
503 (get_casesP ctxt facts @ get_casesT ctxt insts @ [Induct_Args.cases_default])
504 |> tap (trace_rules ctxt casesN)
505 |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
509 |> Seq.maps (Rule_Cases.consume [] facts)
510 |> Seq.maps (fn ((cases, (_, more_facts)), rule) =>
513 |> simp ? (simplified_rule' ctxt #> unmark_constraints);
515 CASES (Rule_Cases.make_common (thy,
516 Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
517 ((Method.insert_tac more_facts THEN' rtac rule' THEN_ALL_NEW
518 (if simp then TRY o trivial_tac else K all_tac)) i) st
526 (** induct method **)
528 val conjunction_congs = [@{thm Pure.all_conjunction}, @{thm imp_conjunction}];
533 fun atomize_term thy =
534 Raw_Simplifier.rewrite_term thy Induct_Args.atomize []
535 #> Object_Logic.drop_judgment thy;
537 val atomize_cterm = Raw_Simplifier.rewrite true Induct_Args.atomize;
539 val atomize_tac = Simplifier.rewrite_goal_tac Induct_Args.atomize;
541 val inner_atomize_tac =
542 Simplifier.rewrite_goal_tac (map Thm.symmetric conjunction_congs) THEN' atomize_tac;
547 fun rulify_term thy =
548 Raw_Simplifier.rewrite_term thy (Induct_Args.rulify @ conjunction_congs) [] #>
549 Raw_Simplifier.rewrite_term thy Induct_Args.rulify_fallback [];
551 fun rulified_term thm =
553 val thy = Thm.theory_of_thm thm;
554 val rulify = rulify_term thy;
555 val (As, B) = Logic.strip_horn (Thm.prop_of thm);
556 in (thy, Logic.list_implies (map rulify As, rulify B)) end;
559 Simplifier.rewrite_goal_tac (Induct_Args.rulify @ conjunction_congs) THEN'
560 Simplifier.rewrite_goal_tac Induct_Args.rulify_fallback THEN'
561 Goal.conjunction_tac THEN_ALL_NEW
562 (Simplifier.rewrite_goal_tac [@{thm Pure.conjunction_imp}] THEN' Goal.norm_hhf_tac);
567 fun rule_instance ctxt inst rule =
568 Drule.cterm_instantiate (prep_inst ctxt align_left I (Thm.prop_of rule, inst)) rule;
570 fun internalize k th =
571 th |> Thm.permute_prems 0 k
572 |> Conv.fconv_rule (Conv.concl_conv (Thm.nprems_of th - k) atomize_cterm);
575 (* guess rule instantiation -- cannot handle pending goal parameters *)
579 fun dest_env thy env =
581 val cert = Thm.cterm_of thy;
582 val certT = Thm.ctyp_of thy;
583 val pairs = Vartab.dest (Envir.term_env env);
584 val types = Vartab.dest (Envir.type_env env);
585 val ts = map (cert o Envir.norm_term env o #2 o #2) pairs;
586 val xs = map2 (curry (cert o Var)) (map #1 pairs) (map (#T o Thm.rep_cterm) ts);
587 in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) types, xs ~~ ts) end;
591 fun guess_instance ctxt rule i st =
593 val thy = Proof_Context.theory_of ctxt;
594 val maxidx = Thm.maxidx_of st;
595 val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*)
596 val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal));
598 if not (null params) then
599 (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^
600 commas_quote (map (Syntax.string_of_term ctxt o Syntax_Trans.mark_boundT) params));
604 val rule' = Thm.incr_indexes (maxidx + 1) rule;
605 val concl = Logic.strip_assums_concl goal;
607 Unify.smash_unifiers thy [(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule'))
608 |> Seq.map (fn env => Drule.instantiate_normalize (dest_env thy env) rule')
611 handle General.Subscript => Seq.empty;
616 (* special renaming of rule parameters *)
618 fun special_rename_params ctxt [[SOME (Free (z, Type (T, _)))]] [thm] =
620 val x = Name.clean (Variable.revert_fixed ctxt z);
622 | index i (y :: ys) =
623 if x = y then x ^ string_of_int i :: index (i + 1) ys
624 else y :: index i ys;
625 fun rename_params [] = []
626 | rename_params ((y, Type (U, _)) :: ys) =
627 (if U = T then x else y) :: rename_params ys
628 | rename_params ((y, _) :: ys) = y :: rename_params ys;
631 val xs = rename_params (Logic.strip_params A);
633 (case filter (fn x' => x' = x) xs of
634 [] => xs | [_] => xs | _ => index 1 xs);
635 in Logic.list_rename_params xs' A end;
637 let val (As, C) = Logic.strip_horn p
638 in Logic.list_implies (map rename_asm As, C) end;
639 val cp' = cterm_fun rename_prop (Thm.cprop_of thm);
640 val thm' = Thm.equal_elim (Thm.reflexive cp') thm;
641 in [Rule_Cases.save thm thm'] end
642 | special_rename_params _ _ ths = ths;
649 fun goal_prefix k ((c as Const ("all", _)) $ Abs (a, T, B)) = c $ Abs (a, T, goal_prefix k B)
650 | goal_prefix 0 _ = Term.dummy_prop
651 | goal_prefix k ((c as Const ("==>", _)) $ A $ B) = c $ A $ goal_prefix (k - 1) B
652 | goal_prefix _ _ = Term.dummy_prop;
654 fun goal_params k (Const ("all", _) $ Abs (_, _, B)) = goal_params k B + 1
655 | goal_params 0 _ = 0
656 | goal_params k (Const ("==>", _) $ _ $ B) = goal_params (k - 1) B
657 | goal_params _ _ = 0;
659 fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) =>
661 val thy = Proof_Context.theory_of ctxt;
662 val cert = Thm.cterm_of thy;
665 fun spec_rule prfx (xs, body) =
666 @{thm Pure.meta_spec}
667 |> Thm.rename_params_rule ([Name.clean (Variable.revert_fixed ctxt x)], 1)
668 |> Thm.lift_rule (cert prfx)
669 |> `(Thm.prop_of #> Logic.strip_assums_concl)
670 |-> (fn pred $ arg =>
671 Drule.cterm_instantiate
672 [(cert (Term.head_of pred), cert (Logic.rlist_abs (xs, body))),
673 (cert (Term.head_of arg), cert (Logic.rlist_abs (xs, v)))]);
675 fun goal_concl k xs (Const ("all", _) $ Abs (a, T, B)) = goal_concl k ((a, T) :: xs) B
676 | goal_concl 0 xs B =
677 if not (Term.exists_subterm (fn t => t aconv v) B) then NONE
678 else SOME (xs, absfree (x, T) (Term.incr_boundvars 1 B))
679 | goal_concl k xs (Const ("==>", _) $ _ $ B) = goal_concl (k - 1) xs B
680 | goal_concl _ _ _ = NONE;
682 (case goal_concl n [] goal of
684 (compose_tac (false, spec_rule (goal_prefix n goal) concl, 1) THEN' rtac asm_rl) i
688 fun miniscope_tac p = CONVERSION o
689 Conv.params_conv p (K (Raw_Simplifier.rewrite true [Thm.symmetric Drule.norm_hhf_eq]));
693 fun arbitrary_tac _ _ [] = K all_tac
694 | arbitrary_tac ctxt n xs = SUBGOAL (fn (goal, i) =>
695 (EVERY' (map (meta_spec_tac ctxt n) xs) THEN'
696 (miniscope_tac (goal_params n goal) ctxt)) i);
703 fun add_defs def_insts =
705 fun add (SOME (_, (t, true))) ctxt = ((SOME t, []), ctxt)
706 | add (SOME (SOME x, (t, _))) ctxt =
707 let val ([(lhs, (_, th))], ctxt') =
708 Local_Defs.add_defs [((x, NoSyn), (Thm.empty_binding, t))] ctxt
709 in ((SOME lhs, [th]), ctxt') end
710 | add (SOME (NONE, (t as Free _, _))) ctxt = ((SOME t, []), ctxt)
711 | add (SOME (NONE, (t, _))) ctxt =
713 val (s, _) = Name.variant "x" (Variable.names_of ctxt);
714 val ([(lhs, (_, th))], ctxt') =
715 Local_Defs.add_defs [((Binding.name s, NoSyn),
716 (Thm.empty_binding, t))] ctxt
717 in ((SOME lhs, [th]), ctxt') end
718 | add NONE ctxt = ((NONE, []), ctxt);
719 in fold_map add def_insts #> apfst (split_list #> apsnd flat) end;
725 rule selection scheme:
726 `A x` induct ... - predicate/set induction
727 induct x - type induction
728 ... induct ... r - explicit rule
731 fun get_inductT ctxt insts =
732 fold_rev (map_product cons) (insts |> map
733 ((fn [] => NONE | ts => List.last ts) #>
734 (fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #>
735 find_inductT ctxt)) [[]]
736 |> filter_out (forall Rule_Cases.is_inner_rule);
738 fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact))
739 | get_inductP _ _ = [];
741 type case_data = (((string * string list) * string list) list * int);
743 fun gen_induct_tac mod_cases ctxt simp def_insts arbitrary taking opt_rule facts =
745 val thy = Proof_Context.theory_of ctxt;
747 val ((insts, defs), defs_ctxt) = fold_map add_defs def_insts ctxt |>> split_list;
748 val atomized_defs = map (map (Conv.fconv_rule atomize_cterm)) defs;
750 fun inst_rule (concls, r) =
751 (if null insts then `Rule_Cases.get r
752 else (align_left "Rule has fewer conclusions than arguments given"
753 (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts
754 |> maps (prep_inst ctxt align_right (atomize_term thy))
755 |> Drule.cterm_instantiate) r |> pair (Rule_Cases.get r))
757 |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th));
761 SOME rs => Seq.single (inst_rule (Rule_Cases.strict_mutual_rule ctxt rs))
763 (get_inductP ctxt facts @
764 map (special_rename_params defs_ctxt insts) (get_inductT ctxt insts))
765 |> map_filter (Rule_Cases.mutual_rule ctxt)
766 |> tap (trace_rules ctxt inductN o map #2)
767 |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
769 fun rule_cases ctxt rule cases =
771 val rule' = Rule_Cases.internalize_params rule;
772 val rule'' = rule' |> simp ? simplified_rule ctxt;
773 val nonames = map (fn ((cn, _), cls) => ((cn, []), cls));
774 val cases' = if Thm.eq_thm_prop (rule', rule'') then cases else nonames cases;
775 in Rule_Cases.make_nested (Thm.prop_of rule'') (rulified_term rule'') cases' end;
779 |> Seq.maps (Rule_Cases.consume (flat defs) facts)
780 |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) =>
781 (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j =>
784 val adefs = nth_list atomized_defs (j - 1);
785 val frees = fold (Term.add_frees o Thm.prop_of) adefs [];
786 val xs = nth_list arbitrary (j - 1);
787 val k = nth concls (j - 1) + more_consumes
789 Method.insert_tac (more_facts @ adefs) THEN'
791 rotate_tac k (length adefs) THEN'
792 arbitrary_tac defs_ctxt k (List.partition (member op = frees) xs |> op @)
794 arbitrary_tac defs_ctxt k xs)
796 THEN' inner_atomize_tac) j))
797 THEN' atomize_tac) i st |> Seq.maps (fn st' =>
798 guess_instance ctxt (internalize more_consumes rule) i st'
799 |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
800 |> Seq.maps (fn rule' =>
801 CASES (rule_cases ctxt rule' cases)
803 PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) st'))))
805 ((if simp then simplify_tac ctxt THEN' (TRY o trivial_tac)
807 THEN_ALL_NEW rulify_tac)
810 val induct_tac = gen_induct_tac (K I);
814 (** coinduct method **)
817 rule selection scheme:
818 goal "A x" coinduct ... - predicate/set coinduction
819 coinduct x - type coinduction
820 coinduct ... r - explicit rule
825 fun get_coinductT ctxt (SOME t :: _) = find_coinductT ctxt (Term.fastype_of t)
826 | get_coinductT _ _ = [];
828 fun get_coinductP ctxt goal = find_coinductP ctxt (Logic.strip_assums_concl goal);
830 fun main_prop_of th =
831 if Rule_Cases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th;
835 fun coinduct_tac ctxt inst taking opt_rule facts =
837 val thy = Proof_Context.theory_of ctxt;
840 if null inst then `Rule_Cases.get r
841 else Drule.cterm_instantiate (prep_inst ctxt align_right I (main_prop_of r, inst)) r
842 |> pair (Rule_Cases.get r);
846 SOME r => Seq.single (inst_rule r)
848 (get_coinductP ctxt goal @ get_coinductT ctxt inst)
849 |> tap (trace_rules ctxt coinductN)
850 |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
852 SUBGOAL_CASES (fn (goal, i) => fn st =>
854 |> Seq.maps (Rule_Cases.consume [] facts)
855 |> Seq.maps (fn ((cases, (_, more_facts)), rule) =>
856 guess_instance ctxt rule i st
857 |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
858 |> Seq.maps (fn rule' =>
859 CASES (Rule_Cases.make_common (thy,
860 Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
861 (Method.insert_tac more_facts i THEN rtac rule' i) st)))
868 (** concrete syntax **)
870 val arbitraryN = "arbitrary";
871 val takingN = "taking";
876 fun single_rule [rule] = rule
877 | single_rule _ = error "Single rule expected";
879 fun named_rule k arg get =
880 Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|--
881 (fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name =>
882 (case get (Context.proof_of context) name of SOME x => x
883 | NONE => error ("No rule for " ^ k ^ " " ^ quote name))))));
885 fun rule get_type get_pred =
886 named_rule typeN (Args.type_name false) get_type ||
887 named_rule predN (Args.const false) get_pred ||
888 named_rule setN (Args.const false) get_pred ||
889 Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms;
891 val cases_rule = rule lookup_casesT lookup_casesP >> single_rule;
892 val induct_rule = rule lookup_inductT lookup_inductP;
893 val coinduct_rule = rule lookup_coinductT lookup_coinductP >> single_rule;
895 val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME;
897 val inst' = Scan.lift (Args.$$$ "_") >> K NONE ||
898 Args.term >> (SOME o rpair false) ||
899 Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --|
900 Scan.lift (Args.$$$ ")");
903 ((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME)
904 -- (Args.term >> rpair false)) >> SOME ||
905 inst' >> Option.map (pair NONE);
907 val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) =>
908 error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
910 fun unless_more_args scan = Scan.unless (Scan.lift
911 ((Args.$$$ arbitraryN || Args.$$$ takingN || Args.$$$ typeN ||
912 Args.$$$ predN || Args.$$$ setN || Args.$$$ ruleN) -- Args.colon)) scan;
914 val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |--
915 Parse.and_list1' (Scan.repeat (unless_more_args free))) [];
917 val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |--
918 Scan.repeat1 (unless_more_args inst)) [];
923 Method.setup @{binding cases}
924 (Args.mode no_simpN --
925 (Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >>
926 (fn (no_simp, (insts, opt_rule)) => fn ctxt =>
927 METHOD_CASES (fn facts => Seq.DETERM (HEADGOAL
928 (cases_tac ctxt (not no_simp) insts opt_rule facts)))))
929 "case analysis on types or predicates/sets";
931 fun gen_induct_setup binding itac =
933 (Args.mode no_simpN -- (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
934 (arbitrary -- taking -- Scan.option induct_rule)) >>
935 (fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn ctxt =>
936 RAW_METHOD_CASES (fn facts =>
938 (HEADGOAL (itac ctxt (not no_simp) insts arbitrary taking opt_rule facts)))))
939 "induction on types or predicates/sets";
941 val induct_setup = gen_induct_setup @{binding induct} induct_tac;
944 Method.setup @{binding coinduct}
945 (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >>
946 (fn ((insts, taking), opt_rule) => fn ctxt =>
947 RAW_METHOD_CASES (fn facts =>
948 Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts)))))
949 "coinduction on types or predicates/sets";
957 val setup = attrib_setup #> cases_setup #> induct_setup #> coinduct_setup;