wenzelm@5824: (* Title: Pure/Isar/method.ML wenzelm@5824: ID: $Id$ wenzelm@5824: Author: Markus Wenzel, TU Muenchen wenzelm@5824: wenzelm@5824: Proof methods. wenzelm@5824: *) wenzelm@5824: wenzelm@5824: signature BASIC_METHOD = wenzelm@5824: sig wenzelm@11731: val trace_rules: bool ref wenzelm@5824: val print_methods: theory -> unit wenzelm@5824: val Method: bstring -> (Args.src -> Proof.context -> Proof.method) -> string -> unit wenzelm@5824: end; wenzelm@5824: wenzelm@5824: signature METHOD = wenzelm@5824: sig wenzelm@5824: include BASIC_METHOD wenzelm@12055: val trace: Proof.context -> thm list -> unit wenzelm@12144: val RAW_METHOD: (thm list -> tactic) -> Proof.method wenzelm@12144: val RAW_METHOD_CASES: wenzelm@12144: (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method wenzelm@6091: val METHOD: (thm list -> tactic) -> Proof.method wenzelm@8372: val METHOD_CASES: wenzelm@8372: (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method wenzelm@9706: val SIMPLE_METHOD: tactic -> Proof.method wenzelm@9706: val SIMPLE_METHOD': ((int -> tactic) -> tactic) -> (int -> tactic) -> Proof.method wenzelm@5824: val fail: Proof.method wenzelm@5824: val succeed: Proof.method wenzelm@8167: val defer: int option -> Proof.method wenzelm@8167: val prefer: int -> Proof.method wenzelm@7419: val insert_tac: thm list -> int -> tactic wenzelm@7574: val insert: thm list -> Proof.method wenzelm@7555: val insert_facts: Proof.method wenzelm@7601: val unfold: thm list -> Proof.method wenzelm@7419: val fold: thm list -> Proof.method wenzelm@7419: val multi_resolve: thm list -> thm -> thm Seq.seq wenzelm@7419: val multi_resolves: thm list -> thm list -> thm Seq.seq wenzelm@12347: val rules_tac: Proof.context -> int option -> int -> tactic wenzelm@6091: val rule_tac: thm list -> thm list -> int -> tactic wenzelm@10309: val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic wenzelm@6091: val rule: thm list -> Proof.method wenzelm@10744: val erule: int -> thm list -> Proof.method wenzelm@10744: val drule: int -> thm list -> Proof.method wenzelm@10744: val frule: int -> thm list -> Proof.method wenzelm@8195: val this: Proof.method wenzelm@7555: val assumption: Proof.context -> Proof.method ballarin@14174: val bires_inst_tac: bool -> Proof.context -> (indexname * string) list -> thm -> int -> tactic wenzelm@8351: val set_tactic: (Proof.context -> thm list -> tactic) -> unit wenzelm@8351: val tactic: string -> Proof.context -> Proof.method wenzelm@5916: exception METHOD_FAIL of (string * Position.T) * exn wenzelm@5824: val method: theory -> Args.src -> Proof.context -> Proof.method wenzelm@9539: val add_method: bstring * (Args.src -> Proof.context -> Proof.method) * string wenzelm@9539: -> theory -> theory wenzelm@5824: val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list wenzelm@5824: -> theory -> theory wenzelm@5884: val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> wenzelm@8282: Args.src -> Proof.context -> Proof.context * 'a wenzelm@8351: val simple_args: (Args.T list -> 'a * Args.T list) wenzelm@8351: -> ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method wenzelm@7555: val ctxt_args: (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method wenzelm@5884: val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method wenzelm@7268: type modifier wenzelm@7601: val sectioned_args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> wenzelm@7268: (Args.T list -> modifier * Args.T list) list -> wenzelm@9864: ('a -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b wenzelm@7601: val bang_sectioned_args: wenzelm@7601: (Args.T list -> modifier * Args.T list) list -> wenzelm@9864: (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a wenzelm@9777: val bang_sectioned_args': wenzelm@9777: (Args.T list -> modifier * Args.T list) list -> wenzelm@9777: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> wenzelm@9864: ('a -> thm list -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b wenzelm@7601: val only_sectioned_args: wenzelm@7601: (Args.T list -> modifier * Args.T list) list -> wenzelm@9864: (Proof.context -> 'a) -> Args.src -> Proof.context -> 'a wenzelm@9864: val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a wenzelm@9864: val thms_args: (thm list -> 'a) -> Args.src -> Proof.context -> 'a wenzelm@9864: val thm_args: (thm -> 'a) -> Args.src -> Proof.context -> 'a wenzelm@5824: datatype text = wenzelm@5824: Basic of (Proof.context -> Proof.method) | wenzelm@5824: Source of Args.src | wenzelm@5824: Then of text list | wenzelm@5824: Orelse of text list | wenzelm@5824: Try of text | wenzelm@5824: Repeat1 of text wenzelm@5824: val refine: text -> Proof.state -> Proof.state Seq.seq wenzelm@8238: val refine_end: text -> Proof.state -> Proof.state Seq.seq wenzelm@5824: val proof: text option -> Proof.state -> Proof.state Seq.seq wenzelm@8966: val local_qed: bool -> text option wenzelm@12144: -> (Proof.context -> string * (string * thm list) list -> unit) * wenzelm@12055: (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq wenzelm@6981: val local_terminal_proof: text * text option wenzelm@12144: -> (Proof.context -> string * (string * thm list) list -> unit) * wenzelm@12055: (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq wenzelm@12144: val local_default_proof: (Proof.context -> string * (string * thm list) list -> unit) * wenzelm@12055: (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq wenzelm@12144: val local_immediate_proof: (Proof.context -> string * (string * thm list) list -> unit) * wenzelm@12055: (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq wenzelm@12144: val local_done_proof: (Proof.context -> string * (string * thm list) list -> unit) * wenzelm@12055: (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq wenzelm@8966: val global_qed: bool -> text option wenzelm@12244: -> Proof.state -> theory * ((string * string) * (string * thm list) list) wenzelm@6934: val global_terminal_proof: text * text option wenzelm@12244: -> Proof.state -> theory * ((string * string) * (string * thm list) list) wenzelm@12244: val global_default_proof: Proof.state -> theory * ((string * string) * (string * thm list) list) wenzelm@12244: val global_immediate_proof: Proof.state -> wenzelm@12244: theory * ((string * string) * (string * thm list) list) wenzelm@12244: val global_done_proof: Proof.state -> theory * ((string * string) * (string * thm list) list) wenzelm@9539: val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic) wenzelm@9539: -> Args.src -> Proof.context -> Proof.method wenzelm@9539: val goal_args': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) wenzelm@9539: -> ('a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method ballarin@14174: val goal_args_ctxt: (Args.T list -> 'a * Args.T list) -> (Proof.context -> 'a -> int -> tactic) ballarin@14174: -> Args.src -> Proof.context -> Proof.method ballarin@14174: val goal_args_ctxt': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ballarin@14174: -> (Proof.context -> 'a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method wenzelm@5824: val setup: (theory -> theory) list wenzelm@5824: end; wenzelm@5824: wenzelm@5824: structure Method: METHOD = wenzelm@5824: struct wenzelm@5824: wenzelm@5824: wenzelm@12324: (** proof methods **) wenzelm@12324: wenzelm@12324: (* tracing *) wenzelm@11731: wenzelm@11731: val trace_rules = ref false; wenzelm@11731: wenzelm@12055: fun trace ctxt rules = wenzelm@12262: conditional (! trace_rules andalso not (null rules)) (fn () => wenzelm@12262: Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules) wenzelm@12262: |> Pretty.string_of |> tracing); wenzelm@11731: wenzelm@11731: wenzelm@8372: (* make methods *) wenzelm@5824: wenzelm@12144: val RAW_METHOD = Proof.method; wenzelm@12144: val RAW_METHOD_CASES = Proof.method_cases; wenzelm@12144: wenzelm@12144: fun METHOD m = Proof.method (fn facts => TRY Tactic.conjunction_tac THEN m facts); wenzelm@12144: fun METHOD_CASES m = wenzelm@12144: Proof.method_cases (fn facts => Seq.THEN (TRY Tactic.conjunction_tac, m facts)); wenzelm@8372: wenzelm@5824: wenzelm@5824: (* primitive *) wenzelm@5824: wenzelm@5824: val fail = METHOD (K no_tac); wenzelm@5824: val succeed = METHOD (K all_tac); wenzelm@5824: wenzelm@5824: wenzelm@8167: (* shuffle *) wenzelm@8167: wenzelm@8240: fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); wenzelm@8167: fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1))); wenzelm@8167: wenzelm@8167: wenzelm@7419: (* insert *) wenzelm@7419: wenzelm@7419: local wenzelm@5824: wenzelm@6981: fun cut_rule_tac raw_rule = wenzelm@6981: let wenzelm@6981: val rule = Drule.forall_intr_vars raw_rule; wenzelm@6981: val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl; wenzelm@7555: in Tactic.rtac (rule COMP revcut_rl) end; wenzelm@6981: wenzelm@7419: in wenzelm@5824: wenzelm@7419: fun insert_tac [] i = all_tac wenzelm@7419: | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); wenzelm@6981: wenzelm@7555: val insert_facts = METHOD (ALLGOALS o insert_tac); wenzelm@7664: fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); wenzelm@7419: wenzelm@9706: fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); wenzelm@9706: fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); wenzelm@9706: wenzelm@12324: end; wenzelm@12324: wenzelm@9706: wenzelm@12384: (* unfold/fold definitions *) wenzelm@12384: wenzelm@12384: fun unfold ths = SIMPLE_METHOD (CHANGED_PROP (rewrite_goals_tac ths)); wenzelm@12384: fun fold ths = SIMPLE_METHOD (CHANGED_PROP (fold_goals_tac ths)); wenzelm@6532: wenzelm@12384: wenzelm@12829: (* atomize rule statements *) wenzelm@12829: wenzelm@12829: fun atomize false = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o ObjectLogic.atomize_tac) wenzelm@12829: | atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac))); wenzelm@12829: wenzelm@12829: wenzelm@12384: (* unfold intro/elim rules *) wenzelm@12384: wenzelm@12384: fun intro ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths)); wenzelm@12384: fun elim ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths)); wenzelm@9484: wenzelm@9484: wenzelm@7419: (* multi_resolve *) wenzelm@7419: wenzelm@7419: local wenzelm@7419: wenzelm@7419: fun res th i rule = wenzelm@7419: Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty; wenzelm@7419: wenzelm@7419: fun multi_res _ [] rule = Seq.single rule wenzelm@7419: | multi_res i (th :: ths) rule = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths rule)); wenzelm@7419: wenzelm@7419: in wenzelm@7419: wenzelm@7419: val multi_resolve = multi_res 1; wenzelm@8372: fun multi_resolves facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules)); wenzelm@7419: wenzelm@7419: end; wenzelm@7419: wenzelm@7419: wenzelm@12347: (* rules_tac *) wenzelm@8372: wenzelm@7419: local wenzelm@5824: wenzelm@12359: val remdups_tac = SUBGOAL (fn (g, i) => wenzelm@12359: let val prems = Logic.strip_assums_hyp g in wenzelm@12359: REPEAT_DETERM_N (length prems - length (gen_distinct op aconv prems)) wenzelm@12359: (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i) wenzelm@12359: end); wenzelm@12347: wenzelm@12347: fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac; wenzelm@12347: wenzelm@12347: fun gen_eq_set e s1 s2 = wenzelm@12347: length s1 = length s2 andalso wenzelm@12347: gen_subset e (s1, s2) andalso gen_subset e (s2, s1); wenzelm@12347: wenzelm@12350: val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist; wenzelm@12347: wenzelm@12347: fun safe_step_tac ctxt = wenzelm@12350: ContextRules.Swrap ctxt wenzelm@12350: (eq_assume_tac ORELSE' wenzelm@12350: bires_tac true (ContextRules.netpair_bang ctxt)); wenzelm@12347: wenzelm@12347: fun unsafe_step_tac ctxt = wenzelm@12350: ContextRules.wrap ctxt wenzelm@12350: (assume_tac APPEND' wenzelm@12350: bires_tac false (ContextRules.netpair_bang ctxt) APPEND' wenzelm@12350: bires_tac false (ContextRules.netpair ctxt)); wenzelm@12347: wenzelm@12347: fun step_tac ctxt i = wenzelm@12347: REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE wenzelm@12347: REMDUPS (unsafe_step_tac ctxt) i; wenzelm@12347: wenzelm@12347: fun intpr_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else wenzelm@12347: let wenzelm@12347: val ps = Logic.strip_assums_hyp g; wenzelm@12347: val c = Logic.strip_assums_concl g; wenzelm@12347: in wenzelm@12347: if gen_mem (fn ((ps1, c1), (ps2, c2)) => wenzelm@12347: c1 aconv c2 andalso gen_eq_set op aconv ps1 ps2) ((ps, c), gs) then no_tac wenzelm@12347: else (step_tac ctxt THEN_ALL_NEW intpr_tac ctxt ((ps, c) :: gs) (d + 1) lim) i wenzelm@12347: end); wenzelm@12347: wenzelm@12347: in wenzelm@12347: wenzelm@12359: fun rules_tac ctxt opt_lim = wenzelm@12359: SELECT_GOAL (DEEPEN (2, if_none opt_lim 20) (intpr_tac ctxt [] 0) 4 1); wenzelm@12347: wenzelm@12347: end; wenzelm@12347: wenzelm@12347: wenzelm@12347: (* rule_tac etc. *) wenzelm@12347: wenzelm@12347: local wenzelm@12347: wenzelm@10541: fun gen_rule_tac tac rules [] i st = tac rules i st wenzelm@12324: | gen_rule_tac tac rules facts i st = wenzelm@12324: Seq.flat (Seq.map (fn rule => (tac o single) rule i st) (multi_resolves facts rules)); wenzelm@7130: wenzelm@10744: fun gen_arule_tac tac j rules facts = wenzelm@10744: EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); wenzelm@10744: wenzelm@11785: fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => wenzelm@11785: let wenzelm@11785: val rules = wenzelm@11785: if not (null arg_rules) then arg_rules wenzelm@12399: else flat (ContextRules.find_rules false facts goal ctxt) wenzelm@12055: in trace ctxt rules; tac rules facts i end); wenzelm@10309: wenzelm@10744: fun meth tac x = METHOD (HEADGOAL o tac x); wenzelm@10744: fun meth' tac x y = METHOD (HEADGOAL o tac x y); wenzelm@8220: wenzelm@7419: in wenzelm@7419: wenzelm@10744: val rule_tac = gen_rule_tac Tactic.resolve_tac; wenzelm@10744: val rule = meth rule_tac; wenzelm@10744: val some_rule_tac = gen_some_rule_tac rule_tac; wenzelm@10744: val some_rule = meth' some_rule_tac; wenzelm@10744: wenzelm@10744: val erule = meth' (gen_arule_tac Tactic.eresolve_tac); wenzelm@10744: val drule = meth' (gen_arule_tac Tactic.dresolve_tac); wenzelm@10744: val frule = meth' (gen_arule_tac Tactic.forward_tac); wenzelm@5824: wenzelm@7419: end; wenzelm@7419: wenzelm@7419: wenzelm@8195: (* this *) wenzelm@8195: wenzelm@8671: val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); wenzelm@8195: wenzelm@8195: wenzelm@8195: (* assumption *) wenzelm@7555: wenzelm@10378: fun asm_tac ths = wenzelm@10378: foldr (op APPEND') (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths, K no_tac); wenzelm@10378: wenzelm@10405: fun assm_tac ctxt = wenzelm@10405: assume_tac APPEND' wenzelm@10405: asm_tac (ProofContext.prems_of ctxt) APPEND' wenzelm@10405: Tactic.rtac Drule.reflexive_thm; wenzelm@7419: wenzelm@7555: fun assumption_tac ctxt [] = assm_tac ctxt wenzelm@10378: | assumption_tac _ [fact] = asm_tac [fact] wenzelm@7555: | assumption_tac _ _ = K no_tac; wenzelm@7419: wenzelm@8671: fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt); wenzelm@7419: wenzelm@7419: wenzelm@9539: (* res_inst_tac etc. *) wenzelm@8238: wenzelm@14718: (*Reimplemented to support both static (Isar) and dynamic (proof state) wenzelm@14718: context. By Clemens Ballarin.*) wenzelm@12119: ballarin@14174: fun bires_inst_tac bires_flag ctxt insts thm = ballarin@14174: let ballarin@14174: val sign = ProofContext.sign_of ctxt; ballarin@14174: (* Separate type and term insts *) ballarin@14174: fun has_type_var ((x, _), _) = (case Symbol.explode x of ballarin@14174: "'"::cs => true | cs => false); ballarin@14174: val Tinsts = filter has_type_var insts; ballarin@14174: val tinsts = filter_out has_type_var insts; ballarin@14174: (* Tactic *) ballarin@14174: fun tac i st = ballarin@14174: let ballarin@14174: (* Preprocess state: extract environment information: ballarin@14174: - variables and their types ballarin@14174: - type variables and their sorts ballarin@14174: - parameters and their types *) ballarin@14174: val (types, sorts) = types_sorts st; ballarin@14174: (* Process type insts: Tinsts_env *) ballarin@14174: fun absent xi = error wenzelm@14718: ("No such variable in theorem: " ^ Syntax.string_of_vname xi); ballarin@14174: val (rtypes, rsorts) = types_sorts thm; ballarin@14174: fun readT (xi, s) = skalberg@15531: let val S = case rsorts xi of SOME S => S | NONE => absent xi; ballarin@14174: val T = Sign.read_typ (sign, sorts) s; ballarin@14174: in if Sign.typ_instance sign (T, TVar (xi, S)) then (xi, T) ballarin@14174: else error ballarin@14174: ("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails") ballarin@14174: end; ballarin@14174: val Tinsts_env = map readT Tinsts; ballarin@14174: (* Preprocess rule: extract vars and their types, apply Tinsts *) ballarin@14174: fun get_typ xi = ballarin@14174: (case rtypes xi of skalberg@15531: SOME T => typ_subst_TVars Tinsts_env T skalberg@15531: | NONE => absent xi); ballarin@14174: val (xis, ss) = Library.split_list tinsts; ballarin@14174: val Ts = map get_typ xis; wenzelm@14718: val (_, _, Bi, _) = dest_state(st,i) wenzelm@14718: val params = Logic.strip_params Bi wenzelm@14718: (* params of subgoal i as string typ pairs *) wenzelm@14718: val params = rev(Term.rename_wrt_term Bi params) wenzelm@14718: (* as they are printed: bound variables with *) ballarin@14508: (* the same name are renamed during printing *) ballarin@14174: fun types' (a, ~1) = (case assoc (params, a) of skalberg@15531: NONE => types (a, ~1) ballarin@14174: | some => some) ballarin@14174: | types' xi = types xi; wenzelm@14718: fun internal x = is_some (types' (x, ~1)); wenzelm@14718: val used = Term.add_term_tvarnames (Thm.prop_of st $ Thm.prop_of thm, []); wenzelm@14718: val (ts, envT) = wenzelm@14718: ProofContext.read_termTs_schematic ctxt internal types' sorts used (ss ~~ Ts); wenzelm@14718: val cenvT = map (apsnd (Thm.ctyp_of sign)) (envT @ Tinsts_env); wenzelm@14718: val cenv = wenzelm@14718: map wenzelm@14718: (fn (xi, t) => wenzelm@14718: pairself (Thm.cterm_of sign) (Var (xi, fastype_of t), t)) wenzelm@14718: (gen_distinct wenzelm@14718: (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) wenzelm@14718: (xis ~~ ts)); wenzelm@14718: (* Lift and instantiate rule *) wenzelm@14718: val {maxidx, ...} = rep_thm st; wenzelm@14718: val paramTs = map #2 params wenzelm@14718: and inc = maxidx+1 wenzelm@14718: fun liftvar (Var ((a,j), T)) = wenzelm@14718: Var((a, j+inc), paramTs ---> incr_tvar inc T) wenzelm@14718: | liftvar t = raise TERM("Variable expected", [t]); wenzelm@14718: fun liftterm t = list_abs_free wenzelm@14718: (params, Logic.incr_indexes(paramTs,inc) t) wenzelm@14718: fun liftpair (cv,ct) = wenzelm@14718: (cterm_fun liftvar cv, cterm_fun liftterm ct) wenzelm@14718: fun lifttvar((a,i),ctyp) = wenzelm@14718: let val {T,sign} = rep_ctyp ctyp wenzelm@14718: in ((a,i+inc), ctyp_of sign (incr_tvar inc T)) end wenzelm@14718: val rule = Drule.instantiate wenzelm@14718: (map lifttvar cenvT, map liftpair cenv) wenzelm@14718: (lift_rule (st, i) thm) ballarin@14174: in wenzelm@14718: if i > nprems_of st then no_tac st wenzelm@14718: else st |> wenzelm@14718: compose_tac (bires_flag, rule, nprems_of thm) i ballarin@14174: end wenzelm@14718: handle TERM (msg,_) => (warning msg; no_tac st) wenzelm@14718: | THM (msg,_,_) => (warning msg; no_tac st); ballarin@14174: in ballarin@14174: tac ballarin@14174: end; wenzelm@8238: ballarin@14174: fun gen_inst _ tac _ (quant, ([], thms)) = ballarin@14174: METHOD (fn facts => quant (insert_tac facts THEN' tac thms)) ballarin@14174: | gen_inst inst_tac _ ctxt (quant, (insts, [thm])) = ballarin@14174: METHOD (fn facts => ballarin@14174: quant (insert_tac facts THEN' inst_tac ctxt insts thm)) ballarin@14174: | gen_inst _ _ _ _ = error "Cannot have instantiations with multiple rules"; wenzelm@14718: ballarin@14174: val res_inst_meth = gen_inst (bires_inst_tac false) Tactic.resolve_tac; ballarin@14174: ballarin@14174: val eres_inst_meth = gen_inst (bires_inst_tac true) Tactic.eresolve_tac; ballarin@14174: ballarin@14174: (* Preserve Var indexes of rl; increment revcut_rl instead. ballarin@14174: Copied from tactic.ML *) ballarin@14174: fun make_elim_preserve rl = ballarin@14174: let val {maxidx,...} = rep_thm rl ballarin@14174: fun cvar xi = cterm_of (Theory.sign_of ProtoPure.thy) (Var(xi,propT)); ballarin@14174: val revcut_rl' = ballarin@14174: instantiate ([], [(cvar("V",0), cvar("V",maxidx+1)), ballarin@14174: (cvar("W",0), cvar("W",maxidx+1))]) revcut_rl ballarin@14174: val arg = (false, rl, nprems_of rl) ballarin@14174: val [th] = Seq.list_of (bicompose false arg 1 revcut_rl') ballarin@14174: in th end ballarin@14174: handle Bind => raise THM("make_elim_preserve", 1, [rl]); wenzelm@8238: ballarin@14174: val cut_inst_meth = ballarin@14174: gen_inst ballarin@14174: (fn ctxt => fn insts => fn thm => ballarin@14174: bires_inst_tac false ctxt insts (make_elim_preserve thm)) ballarin@14174: Tactic.cut_rules_tac; ballarin@14174: ballarin@14174: val dres_inst_meth = ballarin@14174: gen_inst ballarin@14174: (fn ctxt => fn insts => fn rule => ballarin@14174: bires_inst_tac true ctxt insts (make_elim_preserve rule)) ballarin@14174: Tactic.dresolve_tac; ballarin@14174: ballarin@14174: val forw_inst_meth = ballarin@14174: gen_inst ballarin@14174: (fn ctxt => fn insts => fn rule => ballarin@14174: bires_inst_tac false ctxt insts (make_elim_preserve rule) THEN' ballarin@14174: assume_tac) ballarin@14174: Tactic.forward_tac; ballarin@14174: ballarin@14174: fun subgoal_tac ctxt sprop = ballarin@14174: DETERM o bires_inst_tac false ctxt [(("psi", 0), sprop)] cut_rl THEN' ballarin@14174: SUBGOAL (fn (prop, _) => ballarin@14174: let val concl' = Logic.strip_assums_concl prop in ballarin@14174: if null (term_tvars concl') then () ballarin@14174: else warning "Type variables in new subgoal: add a type constraint?"; ballarin@14174: all_tac ballarin@14174: end); ballarin@14174: ballarin@14174: fun subgoals_tac ctxt sprops = EVERY' (map (subgoal_tac ctxt) sprops); ballarin@14174: ballarin@14174: fun thin_tac ctxt s = ballarin@14174: bires_inst_tac true ctxt [(("V", 0), s)] thin_rl; wenzelm@8238: wenzelm@14718: wenzelm@8329: (* simple Prolog interpreter *) wenzelm@8329: wenzelm@8329: fun prolog_tac rules facts = wenzelm@8329: DEPTH_SOLVE_1 (HEADGOAL (Tactic.assume_tac APPEND' Tactic.resolve_tac (facts @ rules))); wenzelm@8329: wenzelm@8329: val prolog = METHOD o prolog_tac; wenzelm@8329: wenzelm@8329: wenzelm@8351: (* ML tactics *) wenzelm@8351: wenzelm@8351: val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic); wenzelm@8351: fun set_tactic f = tactic_ref := f; wenzelm@8351: wenzelm@8351: fun tactic txt ctxt = METHOD (fn facts => wenzelm@9631: (Context.use_mltext wenzelm@9631: ("let fun tactic (ctxt: PureIsar.Proof.context) (facts: thm list) : tactic = \ skalberg@15531: \let val thm = PureIsar.ProofContext.get_thm_closure ctxt o rpair NONE\n\ skalberg@15531: \ and thms = PureIsar.ProofContext.get_thms_closure ctxt o rpair NONE in\n" wenzelm@9631: ^ txt ^ wenzelm@9631: "\nend in PureIsar.Method.set_tactic tactic end") skalberg@15531: false NONE; skalberg@15531: Context.setmp (SOME (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts)); wenzelm@8351: wenzelm@8351: wenzelm@5824: wenzelm@5824: (** methods theory data **) wenzelm@5824: wenzelm@5824: (* data kind 'Isar/methods' *) wenzelm@5824: wenzelm@5824: structure MethodsDataArgs = wenzelm@5824: struct wenzelm@5824: val name = "Isar/methods"; wenzelm@5824: type T = wenzelm@5824: {space: NameSpace.T, wenzelm@5824: meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table}; wenzelm@5824: wenzelm@5824: val empty = {space = NameSpace.empty, meths = Symtab.empty}; wenzelm@6546: val copy = I; wenzelm@5824: val prep_ext = I; wenzelm@5824: fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) = wenzelm@5824: {space = NameSpace.merge (space1, space2), wenzelm@5824: meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups => wenzelm@5824: error ("Attempt to merge different versions of methods " ^ commas_quote dups)}; wenzelm@5824: wenzelm@9222: fun print _ {space, meths} = wenzelm@5824: let wenzelm@5824: fun prt_meth (name, ((_, comment), _)) = Pretty.block wenzelm@6849: [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; wenzelm@5824: in wenzelm@8720: [Pretty.big_list "methods:" (map prt_meth (NameSpace.cond_extern_table space meths))] wenzelm@9222: |> Pretty.chunks |> Pretty.writeln wenzelm@5824: end; wenzelm@5824: end; wenzelm@5824: wenzelm@5824: structure MethodsData = TheoryDataFun(MethodsDataArgs); wenzelm@5824: val print_methods = MethodsData.print; wenzelm@7611: wenzelm@5824: wenzelm@5824: (* get methods *) wenzelm@5824: wenzelm@5916: exception METHOD_FAIL of (string * Position.T) * exn; wenzelm@5916: wenzelm@5824: fun method thy = wenzelm@5824: let wenzelm@5824: val {space, meths} = MethodsData.get thy; wenzelm@5824: wenzelm@5884: fun meth src = wenzelm@5884: let wenzelm@5884: val ((raw_name, _), pos) = Args.dest_src src; wenzelm@5884: val name = NameSpace.intern space raw_name; wenzelm@5884: in wenzelm@5824: (case Symtab.lookup (meths, name) of skalberg@15531: NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos) skalberg@15531: | SOME ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src)) wenzelm@5824: end; wenzelm@5824: in meth end; wenzelm@5824: wenzelm@5824: wenzelm@9194: (* add_method(s) *) wenzelm@5824: wenzelm@5824: fun add_methods raw_meths thy = wenzelm@5824: let wenzelm@5824: val full = Sign.full_name (Theory.sign_of thy); wenzelm@5824: val new_meths = wenzelm@5824: map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths; wenzelm@5824: wenzelm@5824: val {space, meths} = MethodsData.get thy; wenzelm@5824: val space' = NameSpace.extend (space, map fst new_meths); wenzelm@5824: val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups => wenzelm@5824: error ("Duplicate declaration of method(s) " ^ commas_quote dups); wenzelm@5824: in wenzelm@5824: thy |> MethodsData.put {space = space', meths = meths'} wenzelm@5824: end; wenzelm@5824: wenzelm@9194: val add_method = add_methods o Library.single; wenzelm@9194: wenzelm@5824: (*implicit version*) wenzelm@5824: fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]); wenzelm@5824: wenzelm@5824: wenzelm@5884: wenzelm@5884: (** method syntax **) wenzelm@5824: wenzelm@5884: (* basic *) wenzelm@5884: wenzelm@5884: fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) = wenzelm@5884: Args.syntax "method" scan; wenzelm@5824: wenzelm@8351: fun simple_args scan f src ctxt : Proof.method = wenzelm@8351: #2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); wenzelm@8351: wenzelm@7555: fun ctxt_args (f: Proof.context -> Proof.method) src ctxt = wenzelm@8282: #2 (syntax (Scan.succeed (f ctxt)) src ctxt); wenzelm@7555: wenzelm@7555: fun no_args m = ctxt_args (K m); wenzelm@5884: wenzelm@5884: wenzelm@5884: (* sections *) wenzelm@5824: wenzelm@7268: type modifier = (Proof.context -> Proof.context) * Proof.context attribute; wenzelm@7268: wenzelm@7268: local wenzelm@7268: wenzelm@8381: fun sect ss = Scan.first (map Scan.lift ss); wenzelm@5884: fun thms ss = Scan.unless (sect ss) Attrib.local_thms; wenzelm@5884: fun thmss ss = Scan.repeat (thms ss) >> flat; wenzelm@5884: wenzelm@7268: fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]); wenzelm@5824: wenzelm@7268: fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt => wenzelm@7268: Scan.succeed (apply m (ctxt, ths)))) >> #2; wenzelm@5884: wenzelm@7601: fun sectioned args ss = args -- Scan.repeat (section ss); wenzelm@5884: wenzelm@7268: in wenzelm@5824: wenzelm@5884: fun sectioned_args args ss f src ctxt = wenzelm@8282: let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt wenzelm@5921: in f x ctxt' end; wenzelm@5884: wenzelm@7601: fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; wenzelm@9777: fun bang_sectioned_args' ss scan f = wenzelm@9777: sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); wenzelm@7601: fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); wenzelm@7268: wenzelm@8093: fun thms_ctxt_args f = sectioned_args (thmss []) [] f; wenzelm@8093: fun thms_args f = thms_ctxt_args (K o f); wenzelm@9706: fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); wenzelm@5824: wenzelm@7268: end; wenzelm@7268: wenzelm@5824: wenzelm@12347: (* rules syntax *) wenzelm@12347: wenzelm@12347: local wenzelm@12347: wenzelm@12347: val introN = "intro"; wenzelm@12347: val elimN = "elim"; wenzelm@12347: val destN = "dest"; wenzelm@12347: val ruleN = "rule"; wenzelm@12347: wenzelm@12347: fun modifier name kind kind' att = skalberg@15531: Args.$$$ name |-- (kind >> K NONE || kind' |-- Args.nat --| Args.colon >> SOME) wenzelm@12347: >> (pair (I: Proof.context -> Proof.context) o att); wenzelm@12347: wenzelm@12347: val rules_modifiers = wenzelm@12384: [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang_local, wenzelm@12350: modifier destN Args.colon (Scan.succeed ()) ContextRules.dest_local, wenzelm@12350: modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang_local, wenzelm@12350: modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim_local, wenzelm@12350: modifier introN Args.bang_colon Args.bang ContextRules.intro_bang_local, wenzelm@12350: modifier introN Args.colon (Scan.succeed ()) ContextRules.intro_local, wenzelm@12350: Args.del -- Args.colon >> K (I, ContextRules.rule_del_local)]; wenzelm@12347: wenzelm@12347: in wenzelm@12347: wenzelm@12347: fun rules_args m = bang_sectioned_args' rules_modifiers (Scan.lift (Scan.option Args.nat)) m; wenzelm@12347: wenzelm@12347: fun rules_meth n prems ctxt = METHOD (fn facts => wenzelm@12350: HEADGOAL (insert_tac (prems @ facts) THEN' ObjectLogic.atomize_tac THEN' rules_tac ctxt n)); wenzelm@12347: wenzelm@12347: end; wenzelm@12347: wenzelm@12347: wenzelm@9539: (* tactic syntax *) wenzelm@8238: wenzelm@10744: fun nat_thms_args f = uncurry f oo wenzelm@10744: (#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.local_thmss)); wenzelm@10744: wenzelm@8238: val insts = wenzelm@9539: Scan.optional wenzelm@9565: (Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| wenzelm@9565: Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; wenzelm@8238: wenzelm@12119: fun inst_args f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts) src ctxt)); wenzelm@8537: ballarin@14174: val insts_var = ballarin@14174: Scan.optional ballarin@14174: (Args.enum1 "and" (Scan.lift (Args.var -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| ballarin@14174: Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; ballarin@14174: ballarin@14174: fun inst_args_var f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts_var) src ctxt)); wenzelm@8537: wenzelm@12119: fun goal_args' args tac src ctxt = #2 (syntax (Args.goal_spec HEADGOAL -- args >> ballarin@14215: (fn (quant, s) => SIMPLE_METHOD' quant (tac s))) src ctxt); wenzelm@8537: wenzelm@9539: fun goal_args args tac = goal_args' (Scan.lift args) tac; wenzelm@8238: ballarin@14174: fun goal_args_ctxt' args tac src ctxt = ballarin@14174: #2 (syntax (Args.goal_spec HEADGOAL -- args >> ballarin@14174: (fn (quant, s) => SIMPLE_METHOD' quant (tac ctxt s))) src ctxt); wenzelm@8238: ballarin@14174: fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac; wenzelm@5824: wenzelm@14718: wenzelm@5824: (** method text **) wenzelm@5824: wenzelm@5824: (* datatype text *) wenzelm@5824: wenzelm@5824: datatype text = wenzelm@5824: Basic of (Proof.context -> Proof.method) | wenzelm@5824: Source of Args.src | wenzelm@5824: Then of text list | wenzelm@5824: Orelse of text list | wenzelm@5824: Try of text | wenzelm@5824: Repeat1 of text; wenzelm@5824: wenzelm@5824: wenzelm@5824: (* refine *) wenzelm@5824: wenzelm@8238: fun gen_refine f text state = wenzelm@5824: let wenzelm@5824: val thy = Proof.theory_of state; wenzelm@5824: wenzelm@8238: fun eval (Basic mth) = f mth wenzelm@8238: | eval (Source src) = f (method thy src) wenzelm@5824: | eval (Then txts) = Seq.EVERY (map eval txts) wenzelm@5824: | eval (Orelse txts) = Seq.FIRST (map eval txts) wenzelm@5824: | eval (Try txt) = Seq.TRY (eval txt) wenzelm@5824: | eval (Repeat1 txt) = Seq.REPEAT1 (eval txt); wenzelm@5824: in eval text state end; wenzelm@5824: wenzelm@8238: val refine = gen_refine Proof.refine; wenzelm@8238: val refine_end = gen_refine Proof.refine_end; wenzelm@6404: wenzelm@5824: wenzelm@6404: (* structured proof steps *) wenzelm@5824: wenzelm@7506: val default_text = Source (Args.src (("default", []), Position.none)); wenzelm@8195: val this_text = Basic (K this); wenzelm@9706: val done_text = Basic (K (SIMPLE_METHOD all_tac)); wenzelm@7555: wenzelm@8966: fun close_text asm = Basic (fn ctxt => METHOD (K wenzelm@8966: (FILTER Thm.no_prems ((if asm then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)))); wenzelm@8966: skalberg@15531: fun finish_text asm NONE = close_text asm skalberg@15531: | finish_text asm (SOME txt) = Then [txt, close_text asm]; wenzelm@6872: wenzelm@5824: fun proof opt_text state = wenzelm@5824: state wenzelm@5824: |> Proof.assert_backward wenzelm@6404: |> refine (if_none opt_text default_text) wenzelm@8242: |> Seq.map (Proof.goal_facts (K [])) wenzelm@5824: |> Seq.map Proof.enter_forward; wenzelm@5824: wenzelm@8966: fun local_qed asm opt_text = Proof.local_qed (refine (finish_text asm opt_text)); wenzelm@8966: fun local_terminal_proof (text, opt_text) pr = skalberg@15531: Seq.THEN (proof (SOME text), local_qed true opt_text pr); skalberg@15531: val local_default_proof = local_terminal_proof (default_text, NONE); skalberg@15531: val local_immediate_proof = local_terminal_proof (this_text, NONE); skalberg@15531: fun local_done_proof pr = Seq.THEN (proof (SOME done_text), local_qed false NONE pr); wenzelm@5824: wenzelm@6872: wenzelm@8966: fun global_qeds asm opt_text = Proof.global_qed (refine (finish_text asm opt_text)); wenzelm@5824: wenzelm@8966: fun global_qed asm opt_text state = wenzelm@6872: state wenzelm@8966: |> global_qeds asm opt_text wenzelm@6872: |> Proof.check_result "Failed to finish proof" state wenzelm@6872: |> Seq.hd; wenzelm@6872: wenzelm@8966: fun global_term_proof asm (text, opt_text) state = wenzelm@6872: state skalberg@15531: |> proof (SOME text) wenzelm@6872: |> Proof.check_result "Terminal proof method failed" state wenzelm@8966: |> (Seq.flat o Seq.map (global_qeds asm opt_text)) wenzelm@6872: |> Proof.check_result "Failed to finish proof (after successful terminal method)" state wenzelm@6872: |> Seq.hd; wenzelm@6872: wenzelm@8966: val global_terminal_proof = global_term_proof true; skalberg@15531: val global_default_proof = global_terminal_proof (default_text, NONE); skalberg@15531: val global_immediate_proof = global_terminal_proof (this_text, NONE); skalberg@15531: val global_done_proof = global_term_proof false (done_text, NONE); wenzelm@5824: wenzelm@5824: wenzelm@12324: wenzelm@5824: (** theory setup **) wenzelm@5824: wenzelm@9539: (* misc tactic emulations *) wenzelm@9539: ballarin@14174: val subgoal_meth = goal_args_ctxt (Scan.repeat1 Args.name) subgoals_tac; ballarin@14174: val thin_meth = goal_args_ctxt Args.name thin_tac; wenzelm@9539: val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac; wenzelm@9631: val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac; wenzelm@9539: wenzelm@9539: wenzelm@5824: (* pure_methods *) wenzelm@5824: wenzelm@5824: val pure_methods = wenzelm@5824: [("fail", no_args fail, "force failure"), wenzelm@5824: ("succeed", no_args succeed, "succeed"), wenzelm@9587: ("-", no_args insert_facts, "do nothing (insert current facts only)"), wenzelm@9539: ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"), wenzelm@7601: ("unfold", thms_args unfold, "unfold definitions"), wenzelm@12384: ("intro", thms_args intro, "repeatedly apply introduction rules"), wenzelm@12384: ("elim", thms_args elim, "repeatedly apply elimination rules"), wenzelm@7601: ("fold", thms_args fold, "fold definitions"), wenzelm@12829: ("atomize", (atomize o #2) oo syntax (Args.mode "full"), wenzelm@11962: "present local premises as object-level statements"), wenzelm@12347: ("rules", rules_args rules_meth, "apply many rules, including proof search"), wenzelm@12384: ("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"), wenzelm@10744: ("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"), wenzelm@10744: ("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"), wenzelm@10744: ("frule", nat_thms_args frule, "apply rule in forward manner (improper)"), wenzelm@8195: ("this", no_args this, "apply current facts as rules"), wenzelm@8238: ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"), ballarin@14174: ("rule_tac", inst_args_var res_inst_meth, "apply rule (dynamic instantiation)"), ballarin@14174: ("erule_tac", inst_args_var eres_inst_meth, "apply rule in elimination manner (dynamic instantiation)"), ballarin@14174: ("drule_tac", inst_args_var dres_inst_meth, "apply rule in destruct manner (dynamic instantiation)"), ballarin@14174: ("frule_tac", inst_args_var forw_inst_meth, "apply rule in forward manner (dynamic instantiation)"), ballarin@14174: ("cut_tac", inst_args_var cut_inst_meth, "cut rule (dynamic instantiation)"), ballarin@14174: ("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation)"), ballarin@14174: ("thin_tac", thin_meth, "remove premise (dynamic instantiation)"), ballarin@14174: ("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation)"), wenzelm@9631: ("rotate_tac", rotate_meth, "rotate assumptions of goal"), wenzelm@8351: ("prolog", thms_args prolog, "simple prolog interpreter"), wenzelm@8351: ("tactic", simple_args Args.name tactic, "ML tactic as proof method")]; wenzelm@5824: wenzelm@5824: wenzelm@5824: (* setup *) wenzelm@5824: wenzelm@8153: val setup = wenzelm@12324: [MethodsData.init, add_methods pure_methods, skalberg@15531: (#1 o PureThy.add_thms [(("", Drule.equal_intr_rule), [ContextRules.intro_query_global NONE])])]; wenzelm@5824: wenzelm@5824: wenzelm@5824: end; wenzelm@5824: wenzelm@5824: wenzelm@5824: structure BasicMethod: BASIC_METHOD = Method; wenzelm@5824: open BasicMethod;