| author | wenzelm | 
| Mon, 21 Jul 2014 17:37:22 +0200 | |
| changeset 57594 | 037f3b251df5 | 
| parent 55156 | 3ca79ee6eb33 | 
| child 58027 | dc58ab4d9f44 | 
| permissions | -rw-r--r-- | 
| 53707 | 1 | (* Title: Pure/Tools/rule_insts.ML | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 2 | Author: Makarius | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 3 | |
| 53708 | 4 | Rule instantiations -- operations within implicit rule / subgoal context. | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 5 | *) | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 6 | |
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 7 | signature BASIC_RULE_INSTS = | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 8 | sig | 
| 27120 | 9 | val res_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic | 
| 10 | val eres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic | |
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 11 | val cut_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 12 | val forw_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 13 | val dres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 14 | val thin_tac: Proof.context -> string -> int -> tactic | 
| 27219 | 15 | val subgoal_tac: Proof.context -> string -> int -> tactic | 
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 16 | end; | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 17 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 18 | signature RULE_INSTS = | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 19 | sig | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 20 | include BASIC_RULE_INSTS | 
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 21 | val where_rule: Proof.context -> (indexname * string) list -> | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 22 | (binding * string option * mixfix) list -> thm -> thm | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 23 | val of_rule: Proof.context -> string option list * string option list -> | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 24 | (binding * string option * mixfix) list -> thm -> thm | 
| 55151 | 25 | val read_instantiate: Proof.context -> (indexname * string) list -> string list -> thm -> thm | 
| 26 | val instantiate_tac: Proof.context -> (indexname * string) list -> string list -> tactic | |
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 27 | val make_elim_preserve: thm -> thm | 
| 53708 | 28 | val method: (Proof.context -> (indexname * string) list -> thm -> int -> tactic) -> | 
| 29 | (Proof.context -> thm list -> int -> tactic) -> (Proof.context -> Proof.method) context_parser | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 30 | end; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 31 | |
| 42806 | 32 | structure Rule_Insts: RULE_INSTS = | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 33 | struct | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 34 | |
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 35 | (** reading instantiations **) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 36 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 37 | local | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 38 | |
| 45611 | 39 | fun error_var msg xi = error (msg ^ quote (Term.string_of_vname xi)); | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 40 | |
| 45611 | 41 | fun the_sort tvars (xi: indexname) = | 
| 42 | (case AList.lookup (op =) tvars xi of | |
| 43 | SOME S => S | |
| 44 | | NONE => error_var "No such type variable in theorem: " xi); | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 45 | |
| 45611 | 46 | fun the_type vars (xi: indexname) = | 
| 47 | (case AList.lookup (op =) vars xi of | |
| 48 | SOME T => T | |
| 49 | | NONE => error_var "No such variable in theorem: " xi); | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 50 | |
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 51 | fun instantiate inst = | 
| 31977 | 52 | Term_Subst.instantiate ([], map (fn (xi, t) => ((xi, Term.fastype_of t), t)) inst) #> | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 53 | Envir.beta_norm; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 54 | |
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 55 | fun make_instT f v = | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 56 | let | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 57 | val T = TVar v; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 58 | val T' = f T; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 59 | in if T = T' then NONE else SOME (T, T') end; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 60 | |
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 61 | fun make_inst f v = | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 62 | let | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 63 | val t = Var v; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 64 | val t' = f t; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 65 | in if t aconv t' then NONE else SOME (t, t') end; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 66 | |
| 27282 | 67 | val add_used = | 
| 68 | (Thm.fold_terms o fold_types o fold_atyps) | |
| 69 | (fn TFree (a, _) => insert (op =) a | |
| 70 | | TVar ((a, _), _) => insert (op =) a | |
| 71 | | _ => I); | |
| 72 | ||
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 73 | in | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 74 | |
| 55156 | 75 | fun read_termTs ctxt ss Ts = | 
| 25329 | 76 | let | 
| 77 | fun parse T = if T = propT then Syntax.parse_prop ctxt else Syntax.parse_term ctxt; | |
| 78 | val ts = map2 parse Ts ss; | |
| 79 | val ts' = | |
| 39288 | 80 | map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts | 
| 55156 | 81 | |> Syntax.check_terms ctxt | 
| 25329 | 82 | |> Variable.polymorphic ctxt; | 
| 83 | val Ts' = map Term.fastype_of ts'; | |
| 84 | val tyenv = fold Type.raw_match (Ts ~~ Ts') Vartab.empty; | |
| 85 | in (ts', map (apsnd snd) (Vartab.dest tyenv)) end; | |
| 86 | ||
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 87 | fun read_insts ctxt mixed_insts (tvars, vars) = | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 88 | let | 
| 42360 | 89 | val thy = Proof_Context.theory_of ctxt; | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 90 | val cert = Thm.cterm_of thy; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 91 | val certT = Thm.ctyp_of thy; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 92 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 93 | val (type_insts, term_insts) = List.partition (String.isPrefix "'" o fst o fst) mixed_insts; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 94 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 95 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 96 | (* type instantiations *) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 97 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 98 | fun readT (xi, s) = | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 99 | let | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 100 | val S = the_sort tvars xi; | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 101 | val T = Syntax.read_typ ctxt s; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 102 | in | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 103 | if Sign.of_sort thy (T, S) then ((xi, S), T) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 104 | else error_var "Incompatible sort for typ instantiation of " xi | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 105 | end; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 106 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 107 | val instT1 = Term_Subst.instantiateT (map readT type_insts); | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 108 | val vars1 = map (apsnd instT1) vars; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 109 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 110 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 111 | (* term instantiations *) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 112 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 113 | val (xs, ss) = split_list term_insts; | 
| 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 114 | val Ts = map (the_type vars1) xs; | 
| 55156 | 115 | val (ts, inferred) = read_termTs ctxt ss Ts; | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 116 | |
| 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 117 | val instT2 = Term.typ_subst_TVars inferred; | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 118 | val vars2 = map (apsnd instT2) vars1; | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 119 | val inst2 = instantiate (xs ~~ ts); | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 120 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 121 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 122 | (* result *) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 123 | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 124 | val inst_tvars = map_filter (make_instT (instT2 o instT1)) tvars; | 
| 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 125 | val inst_vars = map_filter (make_inst inst2) vars2; | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 126 | in | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 127 | (map (pairself certT) inst_tvars, map (pairself cert) inst_vars) | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 128 | end; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 129 | |
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 130 | fun where_rule ctxt mixed_insts fixes thm = | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 131 | let | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 132 | val ctxt' = ctxt | 
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 133 | |> Proof_Context.read_vars fixes |-> Proof_Context.add_fixes |> #2 | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 134 | |> Variable.declare_thm thm | 
| 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 135 | |> fold (fn a => Variable.declare_names (Logic.mk_type (TFree (a, dummyS)))) (add_used thm []); (* FIXME !? *) | 
| 22692 | 136 | val tvars = Thm.fold_terms Term.add_tvars thm []; | 
| 137 | val vars = Thm.fold_terms Term.add_vars thm []; | |
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 138 | val insts = read_insts ctxt' mixed_insts (tvars, vars); | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 139 | in | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 140 | Drule.instantiate_normalize insts thm | 
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 141 | |> singleton (Proof_Context.export ctxt' ctxt) | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 142 | |> Rule_Cases.save thm | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 143 | end; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 144 | |
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 145 | fun of_rule ctxt (args, concl_args) fixes thm = | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 146 | let | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 147 | fun zip_vars _ [] = [] | 
| 45613 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 148 | | zip_vars (_ :: xs) (NONE :: rest) = zip_vars xs rest | 
| 
70e5b43535cd
simplified read_instantiate -- no longer need to assign values, since rule attributes are now static;
 wenzelm parents: 
45611diff
changeset | 149 | | zip_vars ((x, _) :: xs) (SOME t :: rest) = (x, t) :: zip_vars xs rest | 
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 150 | | zip_vars [] _ = error "More instantiations than variables in theorem"; | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 151 | val insts = | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 152 | zip_vars (rev (Term.add_vars (Thm.full_prop_of thm) [])) args @ | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 153 | zip_vars (rev (Term.add_vars (Thm.concl_of thm) [])) concl_args; | 
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 154 | in where_rule ctxt insts fixes thm end; | 
| 27236 | 155 | |
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 156 | end; | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 157 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 158 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 159 | (* instantiation of rule or goal state *) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 160 | |
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 161 | fun read_instantiate ctxt insts xs = | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 162 | where_rule ctxt insts (map (fn x => (Binding.name x, NONE, NoSyn)) xs); | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 163 | |
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 164 | fun instantiate_tac ctxt insts fixes = | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 165 | PRIMITIVE (read_instantiate ctxt insts fixes); | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 166 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 167 | |
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 168 | |
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 169 | (** attributes **) | 
| 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 170 | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 171 | (* where: named instantiation *) | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 172 | |
| 53171 | 173 | val _ = Theory.setup | 
| 53708 | 174 |   (Attrib.setup @{binding "where"}
 | 
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 175 | (Scan.lift | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 176 | (Parse.and_list (Args.var -- (Args.$$$ "=" |-- Args.name_inner_syntax)) -- Parse.for_fixes) >> | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 177 | (fn (insts, fixes) => | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 178 | Thm.rule_attribute (fn context => where_rule (Context.proof_of context) insts fixes))) | 
| 53171 | 179 | "named instantiation of theorem"); | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 180 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 181 | |
| 20343 
e093a54bf25e
reworked read_instantiate -- separate read_insts;
 wenzelm parents: 
20336diff
changeset | 182 | (* of: positional instantiation (terms only) *) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 183 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 184 | local | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 185 | |
| 55111 | 186 | val inst = Args.maybe Args.name_inner_syntax; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 187 | val concl = Args.$$$ "concl" -- Args.colon; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 188 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 189 | val insts = | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 190 | Scan.repeat (Scan.unless concl inst) -- | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 191 | Scan.optional (concl |-- Scan.repeat inst) []; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 192 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 193 | in | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 194 | |
| 53171 | 195 | val _ = Theory.setup | 
| 53708 | 196 |   (Attrib.setup @{binding "of"}
 | 
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 197 | (Scan.lift (insts -- Parse.for_fixes) >> (fn (args, fixes) => | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
55111diff
changeset | 198 | Thm.rule_attribute (fn context => of_rule (Context.proof_of context) args fixes))) | 
| 53171 | 199 | "positional instantiation of theorem"); | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 200 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 201 | end; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 202 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 203 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 204 | |
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 205 | (** tactics **) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 206 | |
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 207 | (* resolution after lifting and instantation; may refer to parameters of the subgoal *) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 208 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 209 | (* FIXME cleanup this mess!!! *) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 210 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 211 | fun bires_inst_tac bires_flag ctxt insts thm = | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 212 | let | 
| 42360 | 213 | val thy = Proof_Context.theory_of ctxt; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 214 | (* Separate type and term insts *) | 
| 32784 | 215 | fun has_type_var ((x, _), _) = | 
| 216 | (case Symbol.explode x of "'" :: _ => true | _ => false); | |
| 33317 | 217 | val Tinsts = filter has_type_var insts; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 218 | val tinsts = filter_out has_type_var insts; | 
| 25333 | 219 | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 220 | (* Tactic *) | 
| 46474 
7e6be8270ddb
more conventional tactic setup -- avoid low-level Thm.dest_state and spurious warnings about it;
 wenzelm parents: 
46461diff
changeset | 221 | fun tac i st = CSUBGOAL (fn (cgoal, _) => | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 222 | let | 
| 46474 
7e6be8270ddb
more conventional tactic setup -- avoid low-level Thm.dest_state and spurious warnings about it;
 wenzelm parents: 
46461diff
changeset | 223 | val goal = term_of cgoal; | 
| 
7e6be8270ddb
more conventional tactic setup -- avoid low-level Thm.dest_state and spurious warnings about it;
 wenzelm parents: 
46461diff
changeset | 224 | val params = Logic.strip_params goal; (*params of subgoal i as string typ pairs*) | 
| 
7e6be8270ddb
more conventional tactic setup -- avoid low-level Thm.dest_state and spurious warnings about it;
 wenzelm parents: 
46461diff
changeset | 225 | val params = rev (Term.rename_wrt_term goal params) | 
| 25333 | 226 | (*as they are printed: bound variables with*) | 
| 227 | (*the same name are renamed during printing*) | |
| 228 | ||
| 229 | val (param_names, ctxt') = ctxt | |
| 230 | |> Variable.declare_thm thm | |
| 231 | |> Thm.fold_terms Variable.declare_constraints st | |
| 42360 | 232 | |> Proof_Context.add_fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) params); | 
| 25333 | 233 | |
| 234 | (* Process type insts: Tinsts_env *) | |
| 235 | fun absent xi = error | |
| 236 |               ("No such variable in theorem: " ^ Term.string_of_vname xi);
 | |
| 237 | val (rtypes, rsorts) = Drule.types_sorts thm; | |
| 238 | fun readT (xi, s) = | |
| 239 | let val S = case rsorts xi of SOME S => S | NONE => absent xi; | |
| 240 | val T = Syntax.read_typ ctxt' s; | |
| 241 | val U = TVar (xi, S); | |
| 242 | in if Sign.typ_instance thy (T, U) then (U, T) | |
| 243 |                else error ("Instantiation of " ^ Term.string_of_vname xi ^ " fails")
 | |
| 244 | end; | |
| 245 | val Tinsts_env = map readT Tinsts; | |
| 246 | (* Preprocess rule: extract vars and their types, apply Tinsts *) | |
| 247 | fun get_typ xi = | |
| 248 | (case rtypes xi of | |
| 249 | SOME T => typ_subst_atomic Tinsts_env T | |
| 250 | | NONE => absent xi); | |
| 251 | val (xis, ss) = Library.split_list tinsts; | |
| 252 | val Ts = map get_typ xis; | |
| 253 | ||
| 55156 | 254 | val (ts, envT) = | 
| 255 | read_termTs (Proof_Context.set_mode Proof_Context.mode_schematic ctxt') ss Ts; | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 256 | val envT' = map (fn (ixn, T) => | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 257 | (TVar (ixn, the (rsorts ixn)), T)) envT @ Tinsts_env; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 258 | val cenv = | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 259 | map | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 260 | (fn (xi, t) => | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 261 | pairself (Thm.cterm_of thy) (Var (xi, fastype_of t), t)) | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 262 | (distinct | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 263 | (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 264 | (xis ~~ ts)); | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 265 | (* Lift and instantiate rule *) | 
| 44058 | 266 | val maxidx = Thm.maxidx_of st; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 267 | val paramTs = map #2 params | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 268 | and inc = maxidx+1 | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 269 | fun liftvar (Var ((a,j), T)) = | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 270 | Var((a, j+inc), paramTs ---> Logic.incr_tvar inc T) | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 271 |           | liftvar t = raise TERM("Variable expected", [t]);
 | 
| 44241 | 272 | fun liftterm t = | 
| 273 | fold_rev absfree (param_names ~~ paramTs) (Logic.incr_indexes (paramTs, inc) t); | |
| 274 | fun liftpair (cv, ct) = (cterm_fun liftvar cv, cterm_fun liftterm ct); | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 275 | val lifttvar = pairself (ctyp_of thy o Logic.incr_tvar inc); | 
| 43333 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 wenzelm parents: 
42806diff
changeset | 276 | val rule = Drule.instantiate_normalize | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 277 | (map lifttvar envT', map liftpair cenv) | 
| 46474 
7e6be8270ddb
more conventional tactic setup -- avoid low-level Thm.dest_state and spurious warnings about it;
 wenzelm parents: 
46461diff
changeset | 278 | (Thm.lift_rule cgoal thm) | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 279 | in | 
| 46474 
7e6be8270ddb
more conventional tactic setup -- avoid low-level Thm.dest_state and spurious warnings about it;
 wenzelm parents: 
46461diff
changeset | 280 | compose_tac (bires_flag, rule, nprems_of thm) i | 
| 46476 
dac966e4e51d
clarified bires_inst_tac: retain internal exceptions;
 wenzelm parents: 
46474diff
changeset | 281 | end) i st; | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 282 | in tac end; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 283 | |
| 27120 | 284 | val res_inst_tac = bires_inst_tac false; | 
| 285 | val eres_inst_tac = bires_inst_tac true; | |
| 286 | ||
| 287 | ||
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 288 | (* forward resolution *) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 289 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 290 | fun make_elim_preserve rl = | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 291 | let | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 292 | val cert = Thm.cterm_of (Thm.theory_of_thm rl); | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 293 | val maxidx = Thm.maxidx_of rl; | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 294 | fun cvar xi = cert (Var (xi, propT)); | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 295 | val revcut_rl' = | 
| 43333 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 wenzelm parents: 
42806diff
changeset | 296 |       Drule.instantiate_normalize ([], [(cvar ("V", 0), cvar ("V", maxidx + 1)),
 | 
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 297 |         (cvar ("W", 0), cvar ("W", maxidx + 1))]) Drule.revcut_rl;
 | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 298 | in | 
| 52223 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 wenzelm parents: 
46476diff
changeset | 299 | (case Seq.list_of | 
| 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 wenzelm parents: 
46476diff
changeset | 300 |       (Thm.bicompose {flatten = true, match = false, incremented = false}
 | 
| 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 wenzelm parents: 
46476diff
changeset | 301 | (false, rl, Thm.nprems_of rl) 1 revcut_rl') | 
| 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 wenzelm parents: 
46476diff
changeset | 302 | of | 
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 303 | [th] => th | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 304 |     | _ => raise THM ("make_elim_preserve", 1, [rl]))
 | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 305 | end; | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 306 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 307 | (*instantiate and cut -- for atomic fact*) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 308 | fun cut_inst_tac ctxt insts rule = res_inst_tac ctxt insts (make_elim_preserve rule); | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 309 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 310 | (*forward tactic applies a rule to an assumption without deleting it*) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 311 | fun forw_inst_tac ctxt insts rule = cut_inst_tac ctxt insts rule THEN' assume_tac; | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 312 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 313 | (*dresolve tactic applies a rule to replace an assumption*) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 314 | fun dres_inst_tac ctxt insts rule = eres_inst_tac ctxt insts (make_elim_preserve rule); | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 315 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 316 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 317 | (* derived tactics *) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 318 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 319 | (*deletion of an assumption*) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 320 | fun thin_tac ctxt s = eres_inst_tac ctxt [(("V", 0), s)] Drule.thin_rl;
 | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 321 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 322 | (*Introduce the given proposition as lemma and subgoal*) | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 323 | fun subgoal_tac ctxt A = DETERM o res_inst_tac ctxt [(("psi", 0), A)] cut_rl;
 | 
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 324 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 325 | |
| 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 326 | |
| 53708 | 327 | (* method wrapper *) | 
| 27245 
817d34377170
added instantiate_tac, cut_inst_tac, forw_inst_tac, dres_inst_tac, make_elim_preserve (from tactic.ML);
 wenzelm parents: 
27236diff
changeset | 328 | |
| 30545 | 329 | fun method inst_tac tac = | 
| 30515 | 330 | Args.goal_spec -- | 
| 331 | Scan.optional (Scan.lift | |
| 55111 | 332 | (Parse.and_list1 (Args.var -- (Args.$$$ "=" |-- Parse.!!! Args.name_inner_syntax)) --| | 
| 36950 | 333 | Args.$$$ "in")) [] -- | 
| 30515 | 334 | Attrib.thms >> | 
| 335 | (fn ((quant, insts), thms) => fn ctxt => METHOD (fn facts => | |
| 30551 | 336 | if null insts then quant (Method.insert_tac facts THEN' tac ctxt thms) | 
| 30515 | 337 | else | 
| 338 | (case thms of | |
| 339 | [thm] => quant (Method.insert_tac facts THEN' inst_tac ctxt insts thm) | |
| 340 | | _ => error "Cannot have instantiations with multiple rules"))); | |
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 341 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 342 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 343 | (* setup *) | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 344 | |
| 53708 | 345 | (*warning: rule_tac etc. refer to dynamic subgoal context!*) | 
| 346 | ||
| 53171 | 347 | val _ = Theory.setup | 
| 53708 | 348 |  (Method.setup @{binding rule_tac} (method res_inst_tac (K resolve_tac))
 | 
| 349 | "apply rule (dynamic instantiation)" #> | |
| 350 |   Method.setup @{binding erule_tac} (method eres_inst_tac (K eresolve_tac))
 | |
| 30515 | 351 | "apply rule in elimination manner (dynamic instantiation)" #> | 
| 53708 | 352 |   Method.setup @{binding drule_tac} (method dres_inst_tac (K dresolve_tac))
 | 
| 30515 | 353 | "apply rule in destruct manner (dynamic instantiation)" #> | 
| 53708 | 354 |   Method.setup @{binding frule_tac} (method forw_inst_tac (K forward_tac))
 | 
| 30515 | 355 | "apply rule in forward manner (dynamic instantiation)" #> | 
| 53708 | 356 |   Method.setup @{binding cut_tac} (method cut_inst_tac (K cut_rules_tac))
 | 
| 357 | "cut rule (dynamic instantiation)" #> | |
| 358 |   Method.setup @{binding subgoal_tac}
 | |
| 55111 | 359 | (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name_inner_syntax) >> | 
| 46461 | 360 | (fn (quant, props) => fn ctxt => | 
| 361 | SIMPLE_METHOD'' quant (EVERY' (map (subgoal_tac ctxt) props)))) | |
| 30515 | 362 | "insert subgoal (dynamic instantiation)" #> | 
| 53708 | 363 |   Method.setup @{binding thin_tac}
 | 
| 55111 | 364 | (Args.goal_spec -- Scan.lift Args.name_inner_syntax >> | 
| 30515 | 365 | (fn (quant, prop) => fn ctxt => SIMPLE_METHOD'' quant (thin_tac ctxt prop))) | 
| 53171 | 366 | "remove premise (dynamic instantiation)"); | 
| 20336 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 367 | |
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 368 | end; | 
| 
aac494583949
Rule instantiations -- operations within a rule/subgoal context.
 wenzelm parents: diff
changeset | 369 | |
| 42806 | 370 | structure Basic_Rule_Insts: BASIC_RULE_INSTS = Rule_Insts; | 
| 36950 | 371 | open Basic_Rule_Insts; |