wenzelm@24830: (* Title: Tools/induct.ML wenzelm@24830: Author: Markus Wenzel, TU Muenchen wenzelm@24830: wenzelm@26924: Proof by cases, induction, and coinduction. wenzelm@24830: *) wenzelm@24830: wenzelm@37524: signature INDUCT_ARGS = wenzelm@24830: sig wenzelm@24830: val cases_default: thm wenzelm@24830: val atomize: thm list wenzelm@24830: val rulify: thm list wenzelm@24830: val rulify_fallback: thm list berghofe@34987: val equal_def: thm berghofe@34907: val dest_def: term -> (term * term) option wenzelm@58957: val trivial_tac: Proof.context -> int -> tactic wenzelm@24830: end; wenzelm@24830: wenzelm@24830: signature INDUCT = wenzelm@24830: sig wenzelm@24830: (*rule declarations*) wenzelm@24830: val vars_of: term -> term list wenzelm@24830: val dest_rules: Proof.context -> wenzelm@24861: {type_cases: (string * thm) list, pred_cases: (string * thm) list, wenzelm@24861: type_induct: (string * thm) list, pred_induct: (string * thm) list, wenzelm@24861: type_coinduct: (string * thm) list, pred_coinduct: (string * thm) list} wenzelm@24830: val print_rules: Proof.context -> unit wenzelm@24830: val lookup_casesT: Proof.context -> string -> thm option wenzelm@24861: val lookup_casesP: Proof.context -> string -> thm option wenzelm@24830: val lookup_inductT: Proof.context -> string -> thm option wenzelm@24861: val lookup_inductP: Proof.context -> string -> thm option wenzelm@24830: val lookup_coinductT: Proof.context -> string -> thm option wenzelm@24861: val lookup_coinductP: Proof.context -> string -> thm option wenzelm@24830: val find_casesT: Proof.context -> typ -> thm list wenzelm@24861: val find_casesP: Proof.context -> term -> thm list wenzelm@24830: val find_inductT: Proof.context -> typ -> thm list wenzelm@24861: val find_inductP: Proof.context -> term -> thm list wenzelm@24830: val find_coinductT: Proof.context -> typ -> thm list wenzelm@24861: val find_coinductP: Proof.context -> term -> thm list wenzelm@24830: val cases_type: string -> attribute wenzelm@24861: val cases_pred: string -> attribute wenzelm@27140: val cases_del: attribute wenzelm@24830: val induct_type: string -> attribute wenzelm@24861: val induct_pred: string -> attribute wenzelm@27140: val induct_del: attribute wenzelm@24830: val coinduct_type: string -> attribute wenzelm@24861: val coinduct_pred: string -> attribute wenzelm@27140: val coinduct_del: attribute wenzelm@51717: val map_simpset: (Proof.context -> Proof.context) -> Context.generic -> Context.generic wenzelm@36602: val induct_simp_add: attribute wenzelm@36602: val induct_simp_del: attribute berghofe@34907: val no_simpN: string wenzelm@24830: val casesN: string wenzelm@24830: val inductN: string wenzelm@24830: val coinductN: string wenzelm@24830: val typeN: string wenzelm@24861: val predN: string wenzelm@24830: val setN: string wenzelm@24830: (*proof methods*) wenzelm@45132: val arbitrary_tac: Proof.context -> int -> (string * typ) list -> int -> tactic berghofe@34907: val add_defs: (binding option * (term * bool)) option list -> Proof.context -> wenzelm@24830: (term option list * thm list) * Proof.context wenzelm@59970: val atomize_term: Proof.context -> term -> term wenzelm@54742: val atomize_cterm: Proof.context -> conv wenzelm@54742: val atomize_tac: Proof.context -> int -> tactic wenzelm@54742: val inner_atomize_tac: Proof.context -> int -> tactic wenzelm@59970: val rulified_term: Proof.context -> thm -> term wenzelm@54742: val rulify_tac: Proof.context -> int -> tactic berghofe@34907: val simplified_rule: Proof.context -> thm -> thm berghofe@34907: val simplify_tac: Proof.context -> int -> tactic wenzelm@58957: val trivial_tac: Proof.context -> int -> tactic berghofe@34907: val rotate_tac: int -> int -> int -> tactic wenzelm@54742: val internalize: Proof.context -> int -> thm -> thm wenzelm@26940: val guess_instance: Proof.context -> thm -> int -> thm -> thm Seq.seq wenzelm@61844: val cases_context_tactic: bool -> term option list list -> thm option -> wenzelm@61844: thm list -> int -> context_tactic wenzelm@61844: val cases_tac: Proof.context -> bool -> term option list list -> thm option -> wenzelm@61844: thm list -> int -> tactic wenzelm@27323: val get_inductT: Proof.context -> term option list list -> thm list list wenzelm@61844: val gen_induct_context_tactic: ((Rule_Cases.info * int) * thm -> (Rule_Cases.info * int) * thm) -> wenzelm@61841: bool -> (binding option * (term * bool)) option list list -> nipkow@45014: (string * typ) list list -> term option list -> thm list option -> wenzelm@61841: thm list -> int -> context_tactic wenzelm@61844: val gen_induct_tac: Proof.context -> wenzelm@61844: ((Rule_Cases.info * int) * thm -> (Rule_Cases.info * int) * thm) -> wenzelm@61844: bool -> (binding option * (term * bool)) option list list -> wenzelm@61844: (string * typ) list list -> term option list -> thm list option -> wenzelm@61844: thm list -> int -> tactic wenzelm@61844: val induct_context_tactic: bool -> wenzelm@61844: (binding option * (term * bool)) option list list -> wenzelm@26924: (string * typ) list list -> term option list -> thm list option -> wenzelm@61841: thm list -> int -> context_tactic wenzelm@61844: val induct_tac: Proof.context -> bool -> wenzelm@61844: (binding option * (term * bool)) option list list -> wenzelm@61844: (string * typ) list list -> term option list -> thm list option -> wenzelm@61844: thm list -> int -> tactic wenzelm@61844: val coinduct_context_tactic: term option list -> term option list -> thm option -> wenzelm@61841: thm list -> int -> context_tactic wenzelm@61844: val coinduct_tac: Proof.context -> term option list -> term option list -> thm option -> wenzelm@61844: thm list -> int -> tactic nipkow@45014: val gen_induct_setup: binding -> wenzelm@61841: (bool -> (binding option * (term * bool)) option list list -> nipkow@45014: (string * typ) list list -> term option list -> thm list option -> wenzelm@61841: thm list -> int -> context_tactic) -> local_theory -> local_theory wenzelm@24830: end; wenzelm@24830: wenzelm@37524: functor Induct(Induct_Args: INDUCT_ARGS): INDUCT = wenzelm@24830: struct wenzelm@24830: wenzelm@37523: (** variables -- ordered left-to-right, preferring right **) wenzelm@24830: wenzelm@24830: fun vars_of tm = wenzelm@45131: rev (distinct (op =) (Term.fold_aterms (fn t as Var _ => cons t | _ => I) tm [])); wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@37523: val mk_var = Net.encode_type o #2 o Term.dest_Var; wenzelm@24830: wenzelm@47060: fun concl_var which thm = mk_var (which (vars_of (Thm.concl_of thm))) handle List.Empty => wenzelm@24830: raise THM ("No variables in conclusion of rule", 0, [thm]); wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@47060: fun left_var_prem thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle List.Empty => wenzelm@24830: raise THM ("No variables in major premise of rule", 0, [thm]); wenzelm@24830: wenzelm@24830: val left_var_concl = concl_var hd; wenzelm@24830: val right_var_concl = concl_var List.last; wenzelm@24830: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: berghofe@34907: (** constraint simplification **) berghofe@34907: berghofe@34907: (* rearrange parameters and premises to allow application of one-point-rules *) berghofe@34907: berghofe@34907: fun swap_params_conv ctxt i j cv = berghofe@34907: let berghofe@34907: fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt berghofe@34907: | conv1 k ctxt = berghofe@34907: Conv.rewr_conv @{thm swap_params} then_conv wenzelm@60575: Conv.forall_conv (conv1 (k - 1) o snd) ctxt; berghofe@34907: fun conv2 0 ctxt = conv1 j ctxt wenzelm@60575: | conv2 k ctxt = Conv.forall_conv (conv2 (k - 1) o snd) ctxt; berghofe@34907: in conv2 i ctxt end; berghofe@34907: berghofe@34907: fun swap_prems_conv 0 = Conv.all_conv berghofe@34907: | swap_prems_conv i = wenzelm@37525: Conv.implies_concl_conv (swap_prems_conv (i - 1)) then_conv wenzelm@60575: Conv.rewr_conv Drule.swap_prems_eq; berghofe@34907: berghofe@34907: fun find_eq ctxt t = berghofe@34907: let berghofe@34907: val l = length (Logic.strip_params t); berghofe@34907: val Hs = Logic.strip_assums_hyp t; berghofe@34907: fun find (i, t) = wenzelm@59970: (case Induct_Args.dest_def (Object_Logic.drop_judgment ctxt t) of berghofe@34907: SOME (Bound j, _) => SOME (i, j) berghofe@34907: | SOME (_, Bound j) => SOME (i, j) wenzelm@37525: | _ => NONE); berghofe@34907: in wenzelm@37525: (case get_first find (map_index I Hs) of berghofe@34907: NONE => NONE berghofe@34907: | SOME (0, 0) => NONE wenzelm@37525: | SOME (i, j) => SOME (i, l - j - 1, j)) berghofe@34907: end; berghofe@34907: wenzelm@37525: fun mk_swap_rrule ctxt ct = wenzelm@59582: (case find_eq ctxt (Thm.term_of ct) of berghofe@34907: NONE => NONE wenzelm@37525: | SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct)); berghofe@34907: wenzelm@37525: val rearrange_eqs_simproc = wenzelm@69593: Simplifier.make_simproc \<^context> "rearrange_eqs" wenzelm@69593: {lhss = [\<^term>\Pure.all (t :: 'a::{} \ prop)\], wenzelm@62913: proc = fn _ => fn ctxt => fn ct => mk_swap_rrule ctxt ct}; wenzelm@37525: berghofe@34907: berghofe@34907: (* rotate k premises to the left by j, skipping over first j premises *) berghofe@34907: wenzelm@59972: fun rotate_conv 0 _ 0 = Conv.all_conv wenzelm@37525: | rotate_conv 0 j k = swap_prems_conv j then_conv rotate_conv 1 j (k - 1) wenzelm@37525: | rotate_conv i j k = Conv.implies_concl_conv (rotate_conv (i - 1) j k); berghofe@34907: wenzelm@59972: fun rotate_tac _ 0 = K all_tac wenzelm@37525: | rotate_tac j k = SUBGOAL (fn (goal, i) => wenzelm@37525: CONVERSION (rotate_conv wenzelm@37525: j (length (Logic.strip_assums_hyp goal) - j - k) k) i); wenzelm@37525: berghofe@34907: berghofe@34907: (* rulify operators around definition *) berghofe@34907: berghofe@34907: fun rulify_defs_conv ctxt ct = wenzelm@59582: if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) andalso wenzelm@59970: not (is_some (Induct_Args.dest_def (Object_Logic.drop_judgment ctxt (Thm.term_of ct)))) berghofe@34907: then berghofe@34907: (Conv.forall_conv (rulify_defs_conv o snd) ctxt else_conv berghofe@34907: Conv.implies_conv (Conv.try_conv (rulify_defs_conv ctxt)) berghofe@34907: (Conv.try_conv (rulify_defs_conv ctxt)) else_conv wenzelm@37524: Conv.first_conv (map Conv.rewr_conv Induct_Args.rulify) then_conv berghofe@34907: Conv.try_conv (rulify_defs_conv ctxt)) ct berghofe@34907: else Conv.no_conv ct; berghofe@34907: berghofe@34907: berghofe@34907: wenzelm@24830: (** induct data **) wenzelm@24830: wenzelm@24830: (* rules *) wenzelm@24830: wenzelm@30560: type rules = (string * thm) Item_Net.T; wenzelm@24830: wenzelm@33373: fun init_rules index : rules = wenzelm@33373: Item_Net.init wenzelm@33373: (fn ((s1, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm_prop (th1, th2)) wenzelm@33373: (single o index); wenzelm@24830: wenzelm@27140: fun filter_rules (rs: rules) th = wenzelm@30560: filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs); wenzelm@27140: wenzelm@24830: fun pretty_rules ctxt kind rs = wenzelm@30560: let val thms = map snd (Item_Net.content rs) wenzelm@61268: in Pretty.big_list kind (map (Thm.pretty_thm_item ctxt) thms) end; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* context data *) wenzelm@24830: wenzelm@37524: structure Data = Generic_Data wenzelm@24830: ( berghofe@34907: type T = (rules * rules) * (rules * rules) * (rules * rules) * simpset; wenzelm@24830: val empty = wenzelm@24830: ((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)), wenzelm@24830: (init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)), berghofe@34907: (init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2)), wenzelm@69593: simpset_of (empty_simpset \<^context> wenzelm@51717: addsimprocs [rearrange_eqs_simproc] addsimps [Drule.norm_hhf_eq])); wenzelm@24830: val extend = I; berghofe@34907: fun merge (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1), simpset1), berghofe@34907: ((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2), simpset2)) = wenzelm@30560: ((Item_Net.merge (casesT1, casesT2), Item_Net.merge (casesP1, casesP2)), berghofe@34907: (Item_Net.merge (inductT1, inductT2), Item_Net.merge (inductP1, inductP2)), berghofe@34907: (Item_Net.merge (coinductT1, coinductT2), Item_Net.merge (coinductP1, coinductP2)), berghofe@34907: merge_ss (simpset1, simpset2)); wenzelm@24830: ); wenzelm@24830: wenzelm@37524: val get_local = Data.get o Context.Proof; wenzelm@24830: wenzelm@24830: fun dest_rules ctxt = berghofe@34907: let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in wenzelm@30560: {type_cases = Item_Net.content casesT, wenzelm@30560: pred_cases = Item_Net.content casesP, wenzelm@30560: type_induct = Item_Net.content inductT, wenzelm@30560: pred_induct = Item_Net.content inductP, wenzelm@30560: type_coinduct = Item_Net.content coinductT, wenzelm@30560: pred_coinduct = Item_Net.content coinductP} wenzelm@24830: end; wenzelm@24830: wenzelm@24830: fun print_rules ctxt = berghofe@34907: let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in wenzelm@24830: [pretty_rules ctxt "coinduct type:" coinductT, wenzelm@24861: pretty_rules ctxt "coinduct pred:" coinductP, wenzelm@24830: pretty_rules ctxt "induct type:" inductT, wenzelm@24861: pretty_rules ctxt "induct pred:" inductP, wenzelm@24830: pretty_rules ctxt "cases type:" casesT, wenzelm@24861: pretty_rules ctxt "cases pred:" casesP] wenzelm@56334: |> Pretty.writeln_chunks wenzelm@24830: end; wenzelm@24830: wenzelm@24867: val _ = wenzelm@67149: Outer_Syntax.command \<^command_keyword>\print_induct_rules\ wenzelm@46961: "print induction and cases rules" wenzelm@60097: (Scan.succeed (Toplevel.keep (print_rules o Toplevel.context_of))); wenzelm@24830: wenzelm@24830: wenzelm@24830: (* access rules *) wenzelm@24830: wenzelm@61058: local wenzelm@24830: wenzelm@61058: fun lookup_rule which ctxt = wenzelm@61058: AList.lookup (op =) (Item_Net.content (which (get_local ctxt))) wenzelm@67649: #> Option.map (Thm.transfer' ctxt); wenzelm@24830: wenzelm@24830: fun find_rules which how ctxt x = wenzelm@61058: Item_Net.retrieve (which (get_local ctxt)) (how x) wenzelm@67649: |> map (Thm.transfer' ctxt o snd); wenzelm@61058: wenzelm@61058: in wenzelm@61058: wenzelm@61058: val lookup_casesT = lookup_rule (#1 o #1); wenzelm@61058: val lookup_casesP = lookup_rule (#2 o #1); wenzelm@61058: val lookup_inductT = lookup_rule (#1 o #2); wenzelm@61058: val lookup_inductP = lookup_rule (#2 o #2); wenzelm@61058: val lookup_coinductT = lookup_rule (#1 o #3); wenzelm@61058: val lookup_coinductP = lookup_rule (#2 o #3); wenzelm@24830: wenzelm@37523: val find_casesT = find_rules (#1 o #1) Net.encode_type; wenzelm@24861: val find_casesP = find_rules (#2 o #1) I; wenzelm@37523: val find_inductT = find_rules (#1 o #2) Net.encode_type; wenzelm@24861: val find_inductP = find_rules (#2 o #2) I; wenzelm@37523: val find_coinductT = find_rules (#1 o #3) Net.encode_type; wenzelm@24861: val find_coinductP = find_rules (#2 o #3) I; wenzelm@24830: wenzelm@61058: end; wenzelm@61058: wenzelm@24830: wenzelm@24830: wenzelm@24830: (** attributes **) wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@45375: fun mk_att f g name = wenzelm@45375: Thm.mixed_attribute (fn (context, thm) => wenzelm@45375: let wenzelm@45375: val thm' = g thm; wenzelm@61853: val context' = wenzelm@61853: if Thm.is_free_dummy thm then context wenzelm@61853: else Data.map (f (name, Thm.trim_context thm')) context; wenzelm@45375: in (context', thm') end); wenzelm@27140: wenzelm@37525: fun del_att which = wenzelm@59058: Thm.declaration_attribute (fn th => Data.map (which (apply2 (fn rs => wenzelm@37525: fold Item_Net.remove (filter_rules rs th) rs)))); wenzelm@24830: wenzelm@59057: fun add_casesT rule x = @{apply 4(1)} (apfst (Item_Net.update rule)) x; wenzelm@59057: fun add_casesP rule x = @{apply 4(1)} (apsnd (Item_Net.update rule)) x; wenzelm@59057: fun add_inductT rule x = @{apply 4(2)} (apfst (Item_Net.update rule)) x; wenzelm@59057: fun add_inductP rule x = @{apply 4(2)} (apsnd (Item_Net.update rule)) x; wenzelm@59057: fun add_coinductT rule x = @{apply 4(3)} (apfst (Item_Net.update rule)) x; wenzelm@59057: fun add_coinductP rule x = @{apply 4(3)} (apsnd (Item_Net.update rule)) x; wenzelm@24830: wenzelm@45375: val consumes0 = Rule_Cases.default_consumes 0; wenzelm@45375: val consumes1 = Rule_Cases.default_consumes 1; wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@24830: val cases_type = mk_att add_casesT consumes0; wenzelm@24861: val cases_pred = mk_att add_casesP consumes1; wenzelm@59057: val cases_del = del_att @{apply 4(1)}; wenzelm@27140: wenzelm@24830: val induct_type = mk_att add_inductT consumes0; wenzelm@24861: val induct_pred = mk_att add_inductP consumes1; wenzelm@59057: val induct_del = del_att @{apply 4(2)}; wenzelm@27140: wenzelm@24830: val coinduct_type = mk_att add_coinductT consumes0; wenzelm@24861: val coinduct_pred = mk_att add_coinductP consumes1; wenzelm@59057: val coinduct_del = del_att @{apply 4(3)}; wenzelm@24830: wenzelm@51717: fun map_simpset f context = wenzelm@59057: context |> (Data.map o @{apply 4(4)} o Simplifier.simpset_map (Context.proof_of context)) f; wenzelm@36602: wenzelm@36602: fun induct_simp f = wenzelm@51717: Thm.declaration_attribute (fn thm => map_simpset (fn ctxt => f (ctxt, [thm]))); wenzelm@36602: wenzelm@36602: val induct_simp_add = induct_simp (op addsimps); wenzelm@36602: val induct_simp_del = induct_simp (op delsimps); berghofe@34907: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: wenzelm@24830: (** attribute syntax **) wenzelm@24830: berghofe@34907: val no_simpN = "no_simp"; wenzelm@24830: val casesN = "cases"; wenzelm@24830: val inductN = "induct"; wenzelm@24830: val coinductN = "coinduct"; wenzelm@24830: wenzelm@24830: val typeN = "type"; wenzelm@24861: val predN = "pred"; wenzelm@24830: val setN = "set"; wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@24830: fun spec k arg = wenzelm@24830: Scan.lift (Args.$$$ k -- Args.colon) |-- arg || wenzelm@24830: Scan.lift (Args.$$$ k) >> K ""; wenzelm@24830: wenzelm@30528: fun attrib add_type add_pred del = wenzelm@55951: spec typeN (Args.type_name {proper = false, strict = false}) >> add_type || wenzelm@55954: spec predN (Args.const {proper = false, strict = false}) >> add_pred || wenzelm@55954: spec setN (Args.const {proper = false, strict = false}) >> add_pred || wenzelm@30528: Scan.lift Args.del >> K del; wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@58826: val _ = wenzelm@59940: Theory.local_setup wenzelm@67149: (Attrib.local_setup \<^binding>\cases\ (attrib cases_type cases_pred cases_del) wenzelm@58826: "declaration of cases rule" #> wenzelm@67149: Attrib.local_setup \<^binding>\induct\ (attrib induct_type induct_pred induct_del) wenzelm@58826: "declaration of induction rule" #> wenzelm@67149: Attrib.local_setup \<^binding>\coinduct\ (attrib coinduct_type coinduct_pred coinduct_del) wenzelm@58826: "declaration of coinduction rule" #> wenzelm@67149: Attrib.local_setup \<^binding>\induct_simp\ (Attrib.add_del induct_simp_add induct_simp_del) wenzelm@58826: "declaration of rules for simplifying induction or cases rules"); wenzelm@24830: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: wenzelm@24830: (** method utils **) wenzelm@24830: wenzelm@24830: (* alignment *) wenzelm@24830: wenzelm@24830: fun align_left msg xs ys = wenzelm@24830: let val m = length xs and n = length ys haftmann@33957: in if m < n then error msg else (take n xs ~~ ys) end; wenzelm@24830: wenzelm@24830: fun align_right msg xs ys = wenzelm@24830: let val m = length xs and n = length ys haftmann@33957: in if m < n then error msg else (drop (m - n) xs ~~ ys) end; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* prep_inst *) wenzelm@24830: wenzelm@32432: fun prep_inst ctxt align tune (tm, ts) = wenzelm@24830: let wenzelm@60784: fun prep_var (Var (x, xT), SOME t) = wenzelm@24830: let wenzelm@59843: val ct = Thm.cterm_of ctxt (tune t); wenzelm@59586: val tT = Thm.typ_of_cterm ct; wenzelm@24830: in wenzelm@60784: if Type.could_unify (tT, xT) then SOME (x, ct) wenzelm@24830: else error (Pretty.string_of (Pretty.block wenzelm@24830: [Pretty.str "Ill-typed instantiation:", Pretty.fbrk, wenzelm@32432: Syntax.pretty_term ctxt (Thm.term_of ct), Pretty.str " ::", Pretty.brk 1, wenzelm@32432: Syntax.pretty_typ ctxt tT])) wenzelm@24830: end wenzelm@24830: | prep_var (_, NONE) = NONE; wenzelm@24830: val xs = vars_of tm; wenzelm@24830: in wenzelm@24830: align "Rule has fewer variables than instantiations given" xs ts wenzelm@24830: |> map_filter prep_var wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* trace_rules *) wenzelm@24830: wenzelm@24830: fun trace_rules _ kind [] = error ("Unable to figure out " ^ kind ^ " rule") wenzelm@24830: | trace_rules ctxt _ rules = Method.trace ctxt rules; wenzelm@24830: wenzelm@24830: berghofe@34987: (* mark equality constraints in cases rule *) berghofe@34987: wenzelm@37524: val equal_def' = Thm.symmetric Induct_Args.equal_def; berghofe@34987: berghofe@34987: fun mark_constraints n ctxt = Conv.fconv_rule wenzelm@45130: (Conv.prems_conv ~1 (Conv.params_conv ~1 (K (Conv.prems_conv n wenzelm@54742: (Raw_Simplifier.rewrite ctxt false [equal_def']))) ctxt)); berghofe@34987: wenzelm@54742: fun unmark_constraints ctxt = wenzelm@54742: Conv.fconv_rule (Raw_Simplifier.rewrite ctxt true [Induct_Args.equal_def]); berghofe@34987: wenzelm@37525: berghofe@34987: (* simplify *) berghofe@34987: berghofe@34987: fun simplify_conv' ctxt = wenzelm@51717: Simplifier.full_rewrite (put_simpset (#4 (get_local ctxt)) ctxt); berghofe@34987: berghofe@34987: fun simplify_conv ctxt ct = wenzelm@59582: if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) then berghofe@34987: (Conv.try_conv (rulify_defs_conv ctxt) then_conv simplify_conv' ctxt) ct berghofe@34987: else Conv.all_conv ct; berghofe@34987: berghofe@34987: fun gen_simplified_rule cv ctxt = berghofe@34987: Conv.fconv_rule (Conv.prems_conv ~1 (cv ctxt)); berghofe@34987: berghofe@34987: val simplified_rule' = gen_simplified_rule simplify_conv'; berghofe@34987: val simplified_rule = gen_simplified_rule simplify_conv; berghofe@34987: berghofe@34987: fun simplify_tac ctxt = CONVERSION (simplify_conv ctxt); berghofe@34987: wenzelm@37524: val trivial_tac = Induct_Args.trivial_tac; berghofe@34987: berghofe@34987: wenzelm@24830: wenzelm@24830: (** cases method **) wenzelm@24830: wenzelm@24830: (* wenzelm@24830: rule selection scheme: wenzelm@24830: cases - default case split wenzelm@24861: `A t` cases ... - predicate/set cases wenzelm@24830: cases t - type cases wenzelm@24830: ... cases ... r - explicit rule wenzelm@24830: *) wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@24830: fun get_casesT ctxt ((SOME t :: _) :: _) = find_casesT ctxt (Term.fastype_of t) wenzelm@24830: | get_casesT _ _ = []; wenzelm@24830: wenzelm@24861: fun get_casesP ctxt (fact :: _) = find_casesP ctxt (Thm.concl_of fact) wenzelm@24861: | get_casesP _ _ = []; wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@61844: fun cases_context_tactic simp insts opt_rule facts i : context_tactic = wenzelm@61841: fn (ctxt, st) => wenzelm@61841: let wenzelm@61841: fun inst_rule r = wenzelm@61841: (if null insts then r wenzelm@61841: else wenzelm@61841: (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts wenzelm@61841: |> maps (prep_inst ctxt align_left I) wenzelm@61841: |> infer_instantiate ctxt) r) wenzelm@61841: |> simp ? mark_constraints (Rule_Cases.get_constraints r) ctxt wenzelm@61841: |> pair (Rule_Cases.get r); wenzelm@61841: wenzelm@61841: val (opt_rule', facts') = wenzelm@61841: (case (opt_rule, facts) of wenzelm@61841: (NONE, th :: ths) => wenzelm@61841: if is_some (Object_Logic.elim_concl ctxt th) then (SOME th, ths) wenzelm@61841: else (opt_rule, facts) wenzelm@61841: | _ => (opt_rule, facts)); wenzelm@61841: wenzelm@61841: val ruleq = wenzelm@61841: (case opt_rule' of wenzelm@61841: SOME r => Seq.single (inst_rule r) wenzelm@61841: | NONE => wenzelm@61841: (get_casesP ctxt facts' @ get_casesT ctxt insts @ [Induct_Args.cases_default]) wenzelm@61841: |> tap (trace_rules ctxt casesN) wenzelm@61841: |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); wenzelm@61841: in wenzelm@24830: ruleq wenzelm@60455: |> Seq.maps (Rule_Cases.consume ctxt [] facts') wenzelm@60455: |> Seq.maps (fn ((cases, (_, facts'')), rule) => wenzelm@45131: let wenzelm@45131: val rule' = rule wenzelm@54742: |> simp ? (simplified_rule' ctxt #> unmark_constraints ctxt); berghofe@34987: in wenzelm@61841: CONTEXT_CASES (Rule_Cases.make_common ctxt wenzelm@59970: (Thm.prop_of (Rule_Cases.internalize_params rule')) cases) wenzelm@61841: ((Method.insert_tac ctxt facts'' THEN' resolve_tac ctxt [rule'] THEN_ALL_NEW wenzelm@61841: (if simp then TRY o trivial_tac ctxt else K all_tac)) i) (ctxt, st) berghofe@34987: end) wenzelm@61841: end; wenzelm@24830: wenzelm@61844: fun cases_tac ctxt x1 x2 x3 x4 x5 = wenzelm@61844: Method.NO_CONTEXT_TACTIC ctxt (cases_context_tactic x1 x2 x3 x4 x5); wenzelm@61844: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: wenzelm@24830: (** induct method **) wenzelm@24830: wenzelm@59929: val conjunction_congs = @{thms Pure.all_conjunction imp_conjunction}; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* atomize *) wenzelm@24830: wenzelm@59970: fun atomize_term ctxt = wenzelm@59970: Raw_Simplifier.rewrite_term (Proof_Context.theory_of ctxt) Induct_Args.atomize [] wenzelm@59970: #> Object_Logic.drop_judgment ctxt; wenzelm@24830: wenzelm@54742: fun atomize_cterm ctxt = Raw_Simplifier.rewrite ctxt true Induct_Args.atomize; wenzelm@24830: wenzelm@54742: fun atomize_tac ctxt = rewrite_goal_tac ctxt Induct_Args.atomize; wenzelm@24830: wenzelm@54742: fun inner_atomize_tac ctxt = wenzelm@54742: rewrite_goal_tac ctxt (map Thm.symmetric conjunction_congs) THEN' atomize_tac ctxt; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* rulify *) wenzelm@24830: wenzelm@24830: fun rulify_term thy = wenzelm@41228: Raw_Simplifier.rewrite_term thy (Induct_Args.rulify @ conjunction_congs) [] #> wenzelm@41228: Raw_Simplifier.rewrite_term thy Induct_Args.rulify_fallback []; wenzelm@24830: wenzelm@59970: fun rulified_term ctxt thm = wenzelm@24830: let wenzelm@59970: val rulify = rulify_term (Proof_Context.theory_of ctxt); wenzelm@24830: val (As, B) = Logic.strip_horn (Thm.prop_of thm); wenzelm@59970: in Logic.list_implies (map rulify As, rulify B) end; wenzelm@24830: wenzelm@54742: fun rulify_tac ctxt = wenzelm@54742: rewrite_goal_tac ctxt (Induct_Args.rulify @ conjunction_congs) THEN' wenzelm@54742: rewrite_goal_tac ctxt Induct_Args.rulify_fallback THEN' wenzelm@24830: Goal.conjunction_tac THEN_ALL_NEW wenzelm@59929: (rewrite_goal_tac ctxt @{thms Pure.conjunction_imp} THEN' Goal.norm_hhf_tac ctxt); wenzelm@24830: wenzelm@24830: wenzelm@24830: (* prepare rule *) wenzelm@24830: wenzelm@32432: fun rule_instance ctxt inst rule = wenzelm@60784: infer_instantiate ctxt (prep_inst ctxt align_left I (Thm.prop_of rule, inst)) rule; wenzelm@24830: wenzelm@54742: fun internalize ctxt k th = wenzelm@24830: th |> Thm.permute_prems 0 k wenzelm@54742: |> Conv.fconv_rule (Conv.concl_conv (Thm.nprems_of th - k) (atomize_cterm ctxt)); wenzelm@24830: wenzelm@24830: wenzelm@24830: (* guess rule instantiation -- cannot handle pending goal parameters *) wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@59843: fun dest_env ctxt env = wenzelm@24830: let wenzelm@32032: val pairs = Vartab.dest (Envir.term_env env); wenzelm@32032: val types = Vartab.dest (Envir.type_env env); wenzelm@59843: val ts = map (Thm.cterm_of ctxt o Envir.norm_term env o #2 o #2) pairs; wenzelm@60642: val xs = map #1 pairs ~~ map Thm.typ_of_cterm ts; wenzelm@60642: in (map (fn (ai, (S, T)) => ((ai, S), Thm.ctyp_of ctxt T)) types, xs ~~ ts) end; wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@26940: fun guess_instance ctxt rule i st = wenzelm@24830: let wenzelm@26626: val maxidx = Thm.maxidx_of st; wenzelm@24830: val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*) wenzelm@29276: val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal)); wenzelm@24830: in wenzelm@24830: if not (null params) then wenzelm@24830: (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^ wenzelm@49660: commas_quote (map (Syntax.string_of_term ctxt o Syntax_Trans.mark_bound_abs) params)); wenzelm@24830: Seq.single rule) wenzelm@24830: else wenzelm@24830: let wenzelm@24830: val rule' = Thm.incr_indexes (maxidx + 1) rule; wenzelm@24830: val concl = Logic.strip_assums_concl goal; wenzelm@24830: in wenzelm@58950: Unify.smash_unifiers (Context.Proof ctxt) wenzelm@58950: [(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule')) wenzelm@59843: |> Seq.map (fn env => Drule.instantiate_normalize (dest_env ctxt env) rule') wenzelm@24830: end wenzelm@43333: end wenzelm@43333: handle General.Subscript => Seq.empty; wenzelm@24830: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* special renaming of rule parameters *) wenzelm@24830: wenzelm@24830: fun special_rename_params ctxt [[SOME (Free (z, Type (T, _)))]] [thm] = wenzelm@24830: let wenzelm@42488: val x = Name.clean (Variable.revert_fixed ctxt z); wenzelm@59972: fun index _ [] = [] wenzelm@24830: | index i (y :: ys) = wenzelm@24830: if x = y then x ^ string_of_int i :: index (i + 1) ys wenzelm@24830: else y :: index i ys; wenzelm@24830: fun rename_params [] = [] wenzelm@24830: | rename_params ((y, Type (U, _)) :: ys) = wenzelm@24830: (if U = T then x else y) :: rename_params ys wenzelm@24830: | rename_params ((y, _) :: ys) = y :: rename_params ys; wenzelm@24830: fun rename_asm A = wenzelm@24830: let wenzelm@24830: val xs = rename_params (Logic.strip_params A); wenzelm@24830: val xs' = wenzelm@28375: (case filter (fn x' => x' = x) xs of wenzelm@60575: [] => xs wenzelm@60575: | [_] => xs wenzelm@60575: | _ => index 1 xs); wenzelm@45328: in Logic.list_rename_params xs' A end; wenzelm@60313: fun rename_prop prop = wenzelm@60313: let val (As, C) = Logic.strip_horn prop wenzelm@24830: in Logic.list_implies (map rename_asm As, C) end; wenzelm@60313: val thm' = Thm.renamed_prop (rename_prop (Thm.prop_of thm)) thm; wenzelm@33368: in [Rule_Cases.save thm thm'] end wenzelm@24830: | special_rename_params _ _ ths = ths; wenzelm@24830: wenzelm@24830: wenzelm@45132: (* arbitrary_tac *) wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@69593: fun goal_prefix k ((c as Const (\<^const_name>\Pure.all\, _)) $ Abs (a, T, B)) = wenzelm@56245: c $ Abs (a, T, goal_prefix k B) wenzelm@45156: | goal_prefix 0 _ = Term.dummy_prop wenzelm@69593: | goal_prefix k ((c as Const (\<^const_name>\Pure.imp\, _)) $ A $ B) = wenzelm@56245: c $ A $ goal_prefix (k - 1) B wenzelm@45156: | goal_prefix _ _ = Term.dummy_prop; wenzelm@24830: wenzelm@69593: fun goal_params k (Const (\<^const_name>\Pure.all\, _) $ Abs (_, _, B)) = goal_params k B + 1 wenzelm@24830: | goal_params 0 _ = 0 wenzelm@69593: | goal_params k (Const (\<^const_name>\Pure.imp\, _) $ _ $ B) = goal_params (k - 1) B wenzelm@24830: | goal_params _ _ = 0; wenzelm@24830: wenzelm@24830: fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) => wenzelm@24830: let wenzelm@24830: val v = Free (x, T); wenzelm@24830: fun spec_rule prfx (xs, body) = wenzelm@24830: @{thm Pure.meta_spec} wenzelm@42488: |> Thm.rename_params_rule ([Name.clean (Variable.revert_fixed ctxt x)], 1) wenzelm@59843: |> Thm.lift_rule (Thm.cterm_of ctxt prfx) wenzelm@24830: |> `(Thm.prop_of #> Logic.strip_assums_concl) wenzelm@24830: |-> (fn pred $ arg => wenzelm@60784: infer_instantiate ctxt wenzelm@60784: [(#1 (dest_Var (head_of pred)), Thm.cterm_of ctxt (Logic.rlist_abs (xs, body))), wenzelm@60784: (#1 (dest_Var (head_of arg)), Thm.cterm_of ctxt (Logic.rlist_abs (xs, v)))]); wenzelm@24830: wenzelm@69593: fun goal_concl k xs (Const (\<^const_name>\Pure.all\, _) $ Abs (a, T, B)) = wenzelm@56245: goal_concl k ((a, T) :: xs) B wenzelm@24830: | goal_concl 0 xs B = wenzelm@24830: if not (Term.exists_subterm (fn t => t aconv v) B) then NONE wenzelm@44241: else SOME (xs, absfree (x, T) (Term.incr_boundvars 1 B)) wenzelm@69593: | goal_concl k xs (Const (\<^const_name>\Pure.imp\, _) $ _ $ B) = wenzelm@56245: goal_concl (k - 1) xs B wenzelm@24830: | goal_concl _ _ _ = NONE; wenzelm@24830: in wenzelm@24830: (case goal_concl n [] goal of wenzelm@24830: SOME concl => wenzelm@58956: (compose_tac ctxt (false, spec_rule (goal_prefix n goal) concl, 1) THEN' wenzelm@59498: resolve_tac ctxt [asm_rl]) i wenzelm@24830: | NONE => all_tac) wenzelm@24830: end); wenzelm@24830: wenzelm@54742: fun miniscope_tac p = wenzelm@54742: CONVERSION o wenzelm@54742: Conv.params_conv p (fn ctxt => wenzelm@54742: Raw_Simplifier.rewrite ctxt true [Thm.symmetric Drule.norm_hhf_eq]); wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@45132: fun arbitrary_tac _ _ [] = K all_tac wenzelm@45132: | arbitrary_tac ctxt n xs = SUBGOAL (fn (goal, i) => wenzelm@24830: (EVERY' (map (meta_spec_tac ctxt n) xs) THEN' wenzelm@24832: (miniscope_tac (goal_params n goal) ctxt)) i); wenzelm@24830: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* add_defs *) wenzelm@24830: wenzelm@24830: fun add_defs def_insts = wenzelm@24830: let berghofe@34907: fun add (SOME (_, (t, true))) ctxt = ((SOME t, []), ctxt) berghofe@34907: | add (SOME (SOME x, (t, _))) ctxt = wenzelm@28083: let val ([(lhs, (_, th))], ctxt') = wenzelm@63344: Local_Defs.define [((x, NoSyn), ((Thm.def_binding x, []), t))] ctxt wenzelm@24830: in ((SOME lhs, [th]), ctxt') end berghofe@34907: | add (SOME (NONE, (t as Free _, _))) ctxt = ((SOME t, []), ctxt) berghofe@34907: | add (SOME (NONE, (t, _))) ctxt = berghofe@34907: let wenzelm@43326: val (s, _) = Name.variant "x" (Variable.names_of ctxt); wenzelm@49748: val x = Binding.name s; wenzelm@49748: val ([(lhs, (_, th))], ctxt') = ctxt wenzelm@63344: |> Local_Defs.define [((x, NoSyn), ((Thm.def_binding x, []), t))]; berghofe@34907: in ((SOME lhs, [th]), ctxt') end wenzelm@24830: | add NONE ctxt = ((NONE, []), ctxt); wenzelm@24830: in fold_map add def_insts #> apfst (split_list #> apsnd flat) end; wenzelm@24830: wenzelm@24830: wenzelm@24830: (* induct_tac *) wenzelm@24830: wenzelm@24830: (* wenzelm@24830: rule selection scheme: wenzelm@24861: `A x` induct ... - predicate/set induction wenzelm@24830: induct x - type induction wenzelm@24830: ... induct ... r - explicit rule wenzelm@24830: *) wenzelm@24830: wenzelm@24830: fun get_inductT ctxt insts = wenzelm@32188: fold_rev (map_product cons) (insts |> map wenzelm@27323: ((fn [] => NONE | ts => List.last ts) #> wenzelm@27323: (fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #> wenzelm@27323: find_inductT ctxt)) [[]] wenzelm@33368: |> filter_out (forall Rule_Cases.is_inner_rule); wenzelm@24830: wenzelm@24861: fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact)) wenzelm@24861: | get_inductP _ _ = []; wenzelm@24830: wenzelm@61844: fun gen_induct_context_tactic mod_cases simp def_insts arbitrary taking opt_rule facts = wenzelm@61841: CONTEXT_SUBGOAL (fn (_, i) => fn (ctxt, st) => wenzelm@56231: let wenzelm@56231: val ((insts, defs), defs_ctxt) = fold_map add_defs def_insts ctxt |>> split_list; wenzelm@56231: val atomized_defs = map (map (Conv.fconv_rule (atomize_cterm defs_ctxt))) defs; wenzelm@59940: wenzelm@56231: fun inst_rule (concls, r) = wenzelm@56231: (if null insts then `Rule_Cases.get r wenzelm@56231: else (align_left "Rule has fewer conclusions than arguments given" wenzelm@56231: (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts wenzelm@59970: |> maps (prep_inst ctxt align_right (atomize_term ctxt)) wenzelm@60784: |> infer_instantiate ctxt) r |> pair (Rule_Cases.get r)) wenzelm@59970: |> mod_cases wenzelm@56231: |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th)); wenzelm@59940: wenzelm@56231: val ruleq = wenzelm@56231: (case opt_rule of wenzelm@56231: SOME rs => Seq.single (inst_rule (Rule_Cases.strict_mutual_rule ctxt rs)) wenzelm@56231: | NONE => wenzelm@56231: (get_inductP ctxt facts @ wenzelm@56231: map (special_rename_params defs_ctxt insts) (get_inductT ctxt insts)) wenzelm@56231: |> map_filter (Rule_Cases.mutual_rule ctxt) wenzelm@56231: |> tap (trace_rules ctxt inductN o map #2) wenzelm@56231: |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); wenzelm@59940: wenzelm@56231: fun rule_cases ctxt rule cases = wenzelm@56231: let wenzelm@56231: val rule' = Rule_Cases.internalize_params rule; wenzelm@56231: val rule'' = rule' |> simp ? simplified_rule ctxt; wenzelm@56231: val nonames = map (fn ((cn, _), cls) => ((cn, []), cls)); wenzelm@56231: val cases' = if Thm.eq_thm_prop (rule', rule'') then cases else nonames cases; wenzelm@59970: in Rule_Cases.make_nested ctxt (Thm.prop_of rule'') (rulified_term ctxt rule'') cases' end; wenzelm@61841: wenzelm@61841: fun context_tac _ _ = wenzelm@56231: ruleq wenzelm@56231: |> Seq.maps (Rule_Cases.consume defs_ctxt (flat defs) facts) wenzelm@56231: |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => wenzelm@56231: (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => wenzelm@56231: (CONJUNCTS (ALLGOALS wenzelm@56231: let wenzelm@56231: val adefs = nth_list atomized_defs (j - 1); wenzelm@56231: val frees = fold (Term.add_frees o Thm.prop_of) adefs []; wenzelm@56231: val xs = nth_list arbitrary (j - 1); wenzelm@60575: val k = nth concls (j - 1) + more_consumes; wenzelm@56231: in wenzelm@61841: Method.insert_tac defs_ctxt (more_facts @ adefs) THEN' wenzelm@56231: (if simp then wenzelm@56231: rotate_tac k (length adefs) THEN' wenzelm@56231: arbitrary_tac defs_ctxt k (List.partition (member op = frees) xs |> op @) wenzelm@56231: else wenzelm@56231: arbitrary_tac defs_ctxt k xs) wenzelm@56231: end) wenzelm@56231: THEN' inner_atomize_tac defs_ctxt) j)) wenzelm@61841: THEN' atomize_tac defs_ctxt) i st wenzelm@61841: |> Seq.maps (fn st' => wenzelm@56231: guess_instance ctxt (internalize ctxt more_consumes rule) i st' wenzelm@56231: |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking)) wenzelm@56231: |> Seq.maps (fn rule' => wenzelm@61841: CONTEXT_CASES (rule_cases ctxt rule' cases) wenzelm@59498: (resolve_tac ctxt [rule'] i THEN wenzelm@61841: PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) (ctxt, st')))); wenzelm@61841: in wenzelm@61841: (context_tac CONTEXT_THEN_ALL_NEW wenzelm@58957: ((if simp then simplify_tac ctxt THEN' (TRY o trivial_tac ctxt) else K all_tac) wenzelm@61841: THEN_ALL_NEW rulify_tac ctxt)) i (ctxt, st) wenzelm@61841: end) wenzelm@24830: wenzelm@61844: val induct_context_tactic = gen_induct_context_tactic I; wenzelm@61844: wenzelm@61844: fun gen_induct_tac ctxt x1 x2 x3 x4 x5 x6 x7 x8 = wenzelm@61844: Method.NO_CONTEXT_TACTIC ctxt (gen_induct_context_tactic x1 x2 x3 x4 x5 x6 x7 x8); wenzelm@61844: wenzelm@61844: fun induct_tac ctxt x1 x2 x3 x4 x5 x6 x7 = wenzelm@61844: Method.NO_CONTEXT_TACTIC ctxt (induct_context_tactic x1 x2 x3 x4 x5 x6 x7); wenzelm@24830: wenzelm@45130: wenzelm@45130: wenzelm@24830: (** coinduct method **) wenzelm@24830: wenzelm@24830: (* wenzelm@24830: rule selection scheme: wenzelm@24861: goal "A x" coinduct ... - predicate/set coinduction wenzelm@24830: coinduct x - type coinduction wenzelm@24830: coinduct ... r - explicit rule wenzelm@24830: *) wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@24830: fun get_coinductT ctxt (SOME t :: _) = find_coinductT ctxt (Term.fastype_of t) wenzelm@24830: | get_coinductT _ _ = []; wenzelm@24830: wenzelm@24861: fun get_coinductP ctxt goal = find_coinductP ctxt (Logic.strip_assums_concl goal); wenzelm@24861: wenzelm@24861: fun main_prop_of th = wenzelm@33368: if Rule_Cases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th; wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@61844: fun coinduct_context_tactic inst taking opt_rule facts = wenzelm@61841: CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) => wenzelm@61841: let wenzelm@61841: fun inst_rule r = wenzelm@61841: if null inst then `Rule_Cases.get r wenzelm@61841: else wenzelm@61841: infer_instantiate ctxt (prep_inst ctxt align_right I (main_prop_of r, inst)) r wenzelm@61841: |> pair (Rule_Cases.get r); wenzelm@61841: in wenzelm@24830: (case opt_rule of wenzelm@24830: SOME r => Seq.single (inst_rule r) wenzelm@24830: | NONE => wenzelm@24861: (get_coinductP ctxt goal @ get_coinductT ctxt inst) wenzelm@24830: |> tap (trace_rules ctxt coinductN) wenzelm@59972: |> Seq.of_list |> Seq.maps (Seq.try inst_rule)) wenzelm@54742: |> Seq.maps (Rule_Cases.consume ctxt [] facts) wenzelm@24830: |> Seq.maps (fn ((cases, (_, more_facts)), rule) => wenzelm@26940: guess_instance ctxt rule i st wenzelm@32432: |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking)) wenzelm@24830: |> Seq.maps (fn rule' => wenzelm@61841: CONTEXT_CASES (Rule_Cases.make_common ctxt wenzelm@59970: (Thm.prop_of (Rule_Cases.internalize_params rule')) cases) wenzelm@61841: (Method.insert_tac ctxt more_facts i THEN resolve_tac ctxt [rule'] i) (ctxt, st))) wenzelm@61841: end); wenzelm@24830: wenzelm@61844: fun coinduct_tac ctxt x1 x2 x3 x4 x5 = wenzelm@61844: Method.NO_CONTEXT_TACTIC ctxt (coinduct_context_tactic x1 x2 x3 x4 x5); wenzelm@61844: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: wenzelm@24830: wenzelm@24830: (** concrete syntax **) wenzelm@24830: wenzelm@24830: val arbitraryN = "arbitrary"; wenzelm@24830: val takingN = "taking"; wenzelm@24830: val ruleN = "rule"; wenzelm@24830: wenzelm@24830: local wenzelm@24830: wenzelm@24830: fun single_rule [rule] = rule wenzelm@24830: | single_rule _ = error "Single rule expected"; wenzelm@24830: wenzelm@24830: fun named_rule k arg get = wenzelm@24830: Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|-- wenzelm@24830: (fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name => wenzelm@24830: (case get (Context.proof_of context) name of SOME x => x wenzelm@24830: | NONE => error ("No rule for " ^ k ^ " " ^ quote name)))))); wenzelm@24830: wenzelm@24861: fun rule get_type get_pred = wenzelm@55951: named_rule typeN (Args.type_name {proper = false, strict = false}) get_type || wenzelm@55954: named_rule predN (Args.const {proper = false, strict = false}) get_pred || wenzelm@55954: named_rule setN (Args.const {proper = false, strict = false}) get_pred || wenzelm@24830: Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms; wenzelm@24830: wenzelm@24861: val cases_rule = rule lookup_casesT lookup_casesP >> single_rule; wenzelm@24861: val induct_rule = rule lookup_inductT lookup_inductP; wenzelm@24861: val coinduct_rule = rule lookup_coinductT lookup_coinductP >> single_rule; wenzelm@24830: wenzelm@24830: val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME; wenzelm@24830: berghofe@34907: val inst' = Scan.lift (Args.$$$ "_") >> K NONE || berghofe@34907: Args.term >> (SOME o rpair false) || berghofe@34907: Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --| berghofe@34907: Scan.lift (Args.$$$ ")"); berghofe@34907: wenzelm@24830: val def_inst = wenzelm@28083: ((Scan.lift (Args.binding --| (Args.$$$ "\" || Args.$$$ "==")) >> SOME) berghofe@34907: -- (Args.term >> rpair false)) >> SOME || berghofe@34907: inst' >> Option.map (pair NONE); wenzelm@24830: wenzelm@27370: val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => wenzelm@27370: error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); wenzelm@24830: wenzelm@24830: fun unless_more_args scan = Scan.unless (Scan.lift wenzelm@24830: ((Args.$$$ arbitraryN || Args.$$$ takingN || Args.$$$ typeN || wenzelm@24861: Args.$$$ predN || Args.$$$ setN || Args.$$$ ruleN) -- Args.colon)) scan; wenzelm@24830: wenzelm@24830: val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |-- wenzelm@36960: Parse.and_list1' (Scan.repeat (unless_more_args free))) []; wenzelm@24830: wenzelm@24830: val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |-- wenzelm@24830: Scan.repeat1 (unless_more_args inst)) []; wenzelm@24830: wenzelm@24830: in wenzelm@24830: wenzelm@58002: fun gen_induct_setup binding tac = wenzelm@59940: Method.local_setup binding wenzelm@53168: (Scan.lift (Args.mode no_simpN) -- wenzelm@53168: (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) -- wenzelm@53168: (arbitrary -- taking -- Scan.option induct_rule)) >> wenzelm@61841: (fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn _ => fn facts => wenzelm@61841: Seq.DETERM (tac (not no_simp) insts arbitrary taking opt_rule facts 1))) wenzelm@30722: "induction on types or predicates/sets"; wenzelm@24830: wenzelm@58826: val _ = wenzelm@59940: Theory.local_setup wenzelm@67149: (Method.local_setup \<^binding>\cases\ wenzelm@59940: (Scan.lift (Args.mode no_simpN) -- wenzelm@59940: (Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >> wenzelm@61841: (fn (no_simp, (insts, opt_rule)) => fn _ => wenzelm@61841: CONTEXT_METHOD (fn facts => wenzelm@61844: Seq.DETERM (cases_context_tactic (not no_simp) insts opt_rule facts 1)))) wenzelm@59940: "case analysis on types or predicates/sets" #> wenzelm@67149: gen_induct_setup \<^binding>\induct\ induct_context_tactic #> wenzelm@67149: Method.local_setup \<^binding>\coinduct\ wenzelm@58826: (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >> wenzelm@61841: (fn ((insts, taking), opt_rule) => fn _ => fn facts => wenzelm@61844: Seq.DETERM (coinduct_context_tactic insts taking opt_rule facts 1))) wenzelm@58826: "coinduction on types or predicates/sets"); wenzelm@24830: wenzelm@24830: end; wenzelm@24830: wenzelm@24830: end;