| author | wenzelm | 
| Fri, 02 Nov 2001 22:02:41 +0100 | |
| changeset 12021 | 8809efda06d3 | 
| parent 12007 | 72f43e7c3315 | 
| child 12055 | a9c44895cc8c | 
| permissions | -rw-r--r-- | 
| 5824 | 1 | (* Title: Pure/Isar/method.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 8807 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 5824 | 5 | |
| 6 | Proof methods. | |
| 7 | *) | |
| 8 | ||
| 9 | signature BASIC_METHOD = | |
| 10 | sig | |
| 11731 | 11 | val trace_rules: bool ref | 
| 5824 | 12 | val print_methods: theory -> unit | 
| 13 | val Method: bstring -> (Args.src -> Proof.context -> Proof.method) -> string -> unit | |
| 14 | end; | |
| 15 | ||
| 16 | signature METHOD = | |
| 17 | sig | |
| 18 | include BASIC_METHOD | |
| 11731 | 19 | val trace: thm list -> unit | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 20 | val print_global_rules: theory -> unit | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 21 | val print_local_rules: Proof.context -> unit | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 22 | val dest_global: theory attribute | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 23 | val elim_global: theory attribute | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 24 | val intro_global: theory attribute | 
| 11785 | 25 | val dest_bang_global: theory attribute | 
| 26 | val elim_bang_global: theory attribute | |
| 27 | val intro_bang_global: theory attribute | |
| 9938 | 28 | val rule_del_global: theory attribute | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 29 | val dest_local: Proof.context attribute | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 30 | val elim_local: Proof.context attribute | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 31 | val intro_local: Proof.context attribute | 
| 11785 | 32 | val dest_bang_local: Proof.context attribute | 
| 33 | val elim_bang_local: Proof.context attribute | |
| 34 | val intro_bang_local: Proof.context attribute | |
| 9938 | 35 | val rule_del_local: Proof.context attribute | 
| 6091 | 36 | val METHOD: (thm list -> tactic) -> Proof.method | 
| 8372 | 37 | val METHOD_CASES: | 
| 38 | (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method | |
| 9706 | 39 | val SIMPLE_METHOD: tactic -> Proof.method | 
| 40 | val SIMPLE_METHOD': ((int -> tactic) -> tactic) -> (int -> tactic) -> Proof.method | |
| 5824 | 41 | val fail: Proof.method | 
| 42 | val succeed: Proof.method | |
| 8167 | 43 | val defer: int option -> Proof.method | 
| 44 | val prefer: int -> Proof.method | |
| 7419 | 45 | val insert_tac: thm list -> int -> tactic | 
| 7574 | 46 | val insert: thm list -> Proof.method | 
| 7555 | 47 | val insert_facts: Proof.method | 
| 7601 | 48 | val unfold: thm list -> Proof.method | 
| 7419 | 49 | val fold: thm list -> Proof.method | 
| 50 | val multi_resolve: thm list -> thm -> thm Seq.seq | |
| 51 | val multi_resolves: thm list -> thm list -> thm Seq.seq | |
| 6091 | 52 | val rule_tac: thm list -> thm list -> int -> tactic | 
| 10309 | 53 | val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic | 
| 6091 | 54 | val rule: thm list -> Proof.method | 
| 10744 | 55 | val erule: int -> thm list -> Proof.method | 
| 56 | val drule: int -> thm list -> Proof.method | |
| 57 | val frule: int -> thm list -> Proof.method | |
| 8195 | 58 | val this: Proof.method | 
| 7555 | 59 | val assumption: Proof.context -> Proof.method | 
| 8351 | 60 | val set_tactic: (Proof.context -> thm list -> tactic) -> unit | 
| 61 | val tactic: string -> Proof.context -> Proof.method | |
| 5916 | 62 | exception METHOD_FAIL of (string * Position.T) * exn | 
| 5824 | 63 | val method: theory -> Args.src -> Proof.context -> Proof.method | 
| 9539 | 64 | val add_method: bstring * (Args.src -> Proof.context -> Proof.method) * string | 
| 65 | -> theory -> theory | |
| 5824 | 66 | val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list | 
| 67 | -> theory -> theory | |
| 5884 | 68 | val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> | 
| 8282 | 69 | Args.src -> Proof.context -> Proof.context * 'a | 
| 8351 | 70 | val simple_args: (Args.T list -> 'a * Args.T list) | 
| 71 |     -> ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
 | |
| 7555 | 72 | val ctxt_args: (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method | 
| 5884 | 73 | val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method | 
| 7268 | 74 | type modifier | 
| 7601 | 75 | val sectioned_args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> | 
| 7268 | 76 | (Args.T list -> modifier * Args.T list) list -> | 
| 9864 | 77 |     ('a -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
 | 
| 7601 | 78 | val bang_sectioned_args: | 
| 79 | (Args.T list -> modifier * Args.T list) list -> | |
| 9864 | 80 | (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a | 
| 9777 | 81 | val bang_sectioned_args': | 
| 82 | (Args.T list -> modifier * Args.T list) list -> | |
| 83 | (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> | |
| 9864 | 84 |     ('a -> thm list -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
 | 
| 7601 | 85 | val only_sectioned_args: | 
| 86 | (Args.T list -> modifier * Args.T list) list -> | |
| 9864 | 87 | (Proof.context -> 'a) -> Args.src -> Proof.context -> 'a | 
| 88 | val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a | |
| 89 | val thms_args: (thm list -> 'a) -> Args.src -> Proof.context -> 'a | |
| 90 | val thm_args: (thm -> 'a) -> Args.src -> Proof.context -> 'a | |
| 5824 | 91 | datatype text = | 
| 92 | Basic of (Proof.context -> Proof.method) | | |
| 93 | Source of Args.src | | |
| 94 | Then of text list | | |
| 95 | Orelse of text list | | |
| 96 | Try of text | | |
| 97 | Repeat1 of text | |
| 98 | val refine: text -> Proof.state -> Proof.state Seq.seq | |
| 8238 | 99 | val refine_end: text -> Proof.state -> Proof.state Seq.seq | 
| 5824 | 100 | val proof: text option -> Proof.state -> Proof.state Seq.seq | 
| 8966 | 101 | val local_qed: bool -> text option | 
| 6981 | 102 |     -> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
 | 
| 6736 | 103 | -> Proof.state -> Proof.state Seq.seq | 
| 6981 | 104 | val local_terminal_proof: text * text option | 
| 105 |     -> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
 | |
| 6736 | 106 | -> Proof.state -> Proof.state Seq.seq | 
| 6981 | 107 |   val local_default_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
 | 
| 6736 | 108 | -> Proof.state -> Proof.state Seq.seq | 
| 8966 | 109 |   val local_immediate_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
 | 
| 110 | -> Proof.state -> Proof.state Seq.seq | |
| 111 |   val local_done_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
 | |
| 112 | -> Proof.state -> Proof.state Seq.seq | |
| 113 | val global_qed: bool -> text option | |
| 114 |     -> Proof.state -> theory * {kind: string, name: string, thm: thm}
 | |
| 6934 | 115 | val global_terminal_proof: text * text option | 
| 116 |     -> Proof.state -> theory * {kind: string, name: string, thm: thm}
 | |
| 8966 | 117 |   val global_default_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
 | 
| 6532 | 118 |   val global_immediate_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
 | 
| 8966 | 119 |   val global_done_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
 | 
| 9539 | 120 |   val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
 | 
| 121 | -> Args.src -> Proof.context -> Proof.method | |
| 122 | val goal_args': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) | |
| 123 |     -> ('a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method
 | |
| 5824 | 124 | val setup: (theory -> theory) list | 
| 125 | end; | |
| 126 | ||
| 127 | structure Method: METHOD = | |
| 128 | struct | |
| 129 | ||
| 130 | ||
| 11731 | 131 | (** tracing *) | 
| 132 | ||
| 133 | val trace_rules = ref false; | |
| 134 | ||
| 135 | fun trace rules = | |
| 11905 | 136 | if not (! trace_rules) orelse null rules then () | 
| 11731 | 137 | else Pretty.writeln (Pretty.big_list "rules:" (map Display.pretty_thm rules)); | 
| 138 | ||
| 139 | ||
| 140 | ||
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 141 | (** global and local rule data **) | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 142 | |
| 11794 | 143 | val intro_bangK = 0; | 
| 144 | val elim_bangK = 1; | |
| 145 | val introK = 2; | |
| 146 | val elimK = 3; | |
| 11785 | 147 | |
| 10008 | 148 | local | 
| 11785 | 149 | |
| 11794 | 150 | fun kind_name 0 = "safe introduction rules (intro!)" | 
| 151 | | kind_name 1 = "safe elimination rules (elim!)" | |
| 152 | | kind_name 2 = "introduction rules (intro)" | |
| 153 | | kind_name 3 = "elimination rules (elim)" | |
| 154 | | kind_name _ = "unknown"; | |
| 11785 | 155 | |
| 156 | fun prt_rules sg (k, rs) = | |
| 157 | Pretty.writeln (Pretty.big_list (kind_name k ^ ":") | |
| 158 | (map (Display.pretty_thm_sg sg) (NetRules.rules rs))); | |
| 159 | ||
| 10008 | 160 | in | 
| 11785 | 161 | fun print_rules sg rules = seq (prt_rules sg) (0 upto length rules - 1 ~~ rules); | 
| 10008 | 162 | end; | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 163 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 164 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 165 | (* theory data kind 'Isar/rules' *) | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 166 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 167 | structure GlobalRulesArgs = | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 168 | struct | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 169 | val name = "Isar/rules"; | 
| 11785 | 170 | type T = thm NetRules.T list; | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 171 | |
| 11785 | 172 | val empty = [NetRules.intro, NetRules.elim, NetRules.intro, NetRules.elim]; | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 173 | val copy = I; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 174 | val prep_ext = I; | 
| 11785 | 175 | fun merge x = map2 NetRules.merge x; | 
| 10008 | 176 | val print = print_rules; | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 177 | end; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 178 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 179 | structure GlobalRules = TheoryDataFun(GlobalRulesArgs); | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 180 | val print_global_rules = GlobalRules.print; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 181 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 182 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 183 | (* proof data kind 'Isar/rules' *) | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 184 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 185 | structure LocalRulesArgs = | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 186 | struct | 
| 11785 | 187 | val name = GlobalRulesArgs.name; | 
| 188 | type T = GlobalRulesArgs.T; | |
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 189 | val init = GlobalRules.get; | 
| 10008 | 190 | val print = print_rules o ProofContext.sign_of; | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 191 | end; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 192 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 193 | structure LocalRules = ProofDataFun(LocalRulesArgs); | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 194 | val print_local_rules = LocalRules.print; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 195 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 196 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 197 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 198 | (** attributes **) | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 199 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 200 | (* add rules *) | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 201 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 202 | local | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 203 | |
| 11785 | 204 | fun add_rule k view th = Library.map_nth_elem k (NetRules.insert (view th)); | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 205 | |
| 11785 | 206 | val add_intro = add_rule introK I; | 
| 207 | val add_elim = add_rule elimK I; | |
| 208 | val add_dest = add_rule elimK Tactic.make_elim; | |
| 209 | val add_intro_bang = add_rule intro_bangK I; | |
| 210 | val add_elim_bang = add_rule elim_bangK I; | |
| 211 | val add_dest_bang = add_rule elim_bangK Tactic.make_elim; | |
| 212 | ||
| 213 | fun del_rule th = map (NetRules.delete th o NetRules.delete (Tactic.make_elim th)); | |
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 214 | |
| 9938 | 215 | fun mk_att f g (x, th) = (f (g th) x, th); | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 216 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 217 | in | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 218 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 219 | val dest_global = mk_att GlobalRules.map add_dest; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 220 | val elim_global = mk_att GlobalRules.map add_elim; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 221 | val intro_global = mk_att GlobalRules.map add_intro; | 
| 11785 | 222 | val dest_bang_global = mk_att GlobalRules.map add_dest_bang; | 
| 223 | val elim_bang_global = mk_att GlobalRules.map add_elim_bang; | |
| 224 | val intro_bang_global = mk_att GlobalRules.map add_intro_bang; | |
| 9938 | 225 | val rule_del_global = mk_att GlobalRules.map del_rule; | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 226 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 227 | val dest_local = mk_att LocalRules.map add_dest; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 228 | val elim_local = mk_att LocalRules.map add_elim; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 229 | val intro_local = mk_att LocalRules.map add_intro; | 
| 11785 | 230 | val dest_bang_local = mk_att LocalRules.map add_dest_bang; | 
| 231 | val elim_bang_local = mk_att LocalRules.map add_elim_bang; | |
| 232 | val intro_bang_local = mk_att LocalRules.map add_intro_bang; | |
| 9938 | 233 | val rule_del_local = mk_att LocalRules.map del_rule; | 
| 234 | ||
| 10034 | 235 | fun del_args att = Attrib.syntax (Scan.lift Args.del >> K att); | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 236 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 237 | end; | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 238 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 239 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 240 | (* concrete syntax *) | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 241 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 242 | val rule_atts = | 
| 11785 | 243 |  [("dest",
 | 
| 244 | (Attrib.bang_args dest_bang_global dest_global, | |
| 245 | Attrib.bang_args dest_bang_local dest_local), | |
| 9900 | 246 | "declaration of destruction rule"), | 
| 11785 | 247 |   ("elim",
 | 
| 248 | (Attrib.bang_args elim_bang_global elim_global, | |
| 249 | Attrib.bang_args elim_bang_local elim_local), | |
| 9900 | 250 | "declaration of elimination rule"), | 
| 11785 | 251 |   ("intro",
 | 
| 252 | (Attrib.bang_args intro_bang_global intro_global, | |
| 253 | Attrib.bang_args intro_bang_local intro_local), | |
| 9900 | 254 | "declaration of introduction rule"), | 
| 9938 | 255 |   ("rule", (del_args rule_del_global, del_args rule_del_local),
 | 
| 9900 | 256 | "remove declaration of elim/intro rule")]; | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 257 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 258 | |
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 259 | |
| 5824 | 260 | (** proof methods **) | 
| 261 | ||
| 8372 | 262 | (* make methods *) | 
| 5824 | 263 | |
| 6849 | 264 | val METHOD = Proof.method; | 
| 8372 | 265 | val METHOD_CASES = Proof.method_cases; | 
| 266 | ||
| 5824 | 267 | |
| 268 | (* primitive *) | |
| 269 | ||
| 270 | val fail = METHOD (K no_tac); | |
| 271 | val succeed = METHOD (K all_tac); | |
| 272 | ||
| 273 | ||
| 8167 | 274 | (* shuffle *) | 
| 275 | ||
| 8240 | 276 | fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); | 
| 8167 | 277 | fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1))); | 
| 278 | ||
| 279 | ||
| 7419 | 280 | (* insert *) | 
| 281 | ||
| 282 | local | |
| 5824 | 283 | |
| 6981 | 284 | fun cut_rule_tac raw_rule = | 
| 285 | let | |
| 286 | val rule = Drule.forall_intr_vars raw_rule; | |
| 287 | val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl; | |
| 7555 | 288 | in Tactic.rtac (rule COMP revcut_rl) end; | 
| 6981 | 289 | |
| 7419 | 290 | in | 
| 5824 | 291 | |
| 7419 | 292 | fun insert_tac [] i = all_tac | 
| 293 | | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); | |
| 6981 | 294 | |
| 7555 | 295 | val insert_facts = METHOD (ALLGOALS o insert_tac); | 
| 7664 | 296 | fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); | 
| 7419 | 297 | |
| 298 | end; | |
| 5824 | 299 | |
| 300 | ||
| 9706 | 301 | (* simple methods *) | 
| 302 | ||
| 303 | fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); | |
| 304 | fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); | |
| 305 | ||
| 306 | ||
| 7601 | 307 | (* unfold / fold definitions *) | 
| 6532 | 308 | |
| 10821 | 309 | fun unfold thms = SIMPLE_METHOD (CHANGED_PROP (rewrite_goals_tac thms)); | 
| 310 | fun fold thms = SIMPLE_METHOD (CHANGED_PROP (fold_goals_tac thms)); | |
| 9484 | 311 | |
| 312 | ||
| 7419 | 313 | (* multi_resolve *) | 
| 314 | ||
| 315 | local | |
| 316 | ||
| 317 | fun res th i rule = | |
| 318 | Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty; | |
| 319 | ||
| 320 | fun multi_res _ [] rule = Seq.single rule | |
| 321 | | multi_res i (th :: ths) rule = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths rule)); | |
| 322 | ||
| 323 | in | |
| 324 | ||
| 325 | val multi_resolve = multi_res 1; | |
| 8372 | 326 | fun multi_resolves facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules)); | 
| 7419 | 327 | |
| 328 | end; | |
| 329 | ||
| 330 | ||
| 10744 | 331 | (* basic rules *) | 
| 8372 | 332 | |
| 7419 | 333 | local | 
| 5824 | 334 | |
| 10541 
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
 wenzelm parents: 
10529diff
changeset | 335 | fun gen_rule_tac tac rules [] i st = tac rules i st | 
| 
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
 wenzelm parents: 
10529diff
changeset | 336 | | gen_rule_tac tac erules facts i st = | 
| 
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
 wenzelm parents: 
10529diff
changeset | 337 | Seq.flat (Seq.map (fn rule => (tac o single) rule i st) (multi_resolves facts erules)); | 
| 7130 | 338 | |
| 10744 | 339 | fun gen_arule_tac tac j rules facts = | 
| 340 | EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); | |
| 341 | ||
| 11785 | 342 | fun find_erules [] _ = [] | 
| 343 | | find_erules (fact :: _) rs = | |
| 344 | NetRules.retrieve rs (Logic.strip_assums_concl (#prop (Thm.rep_thm fact))); | |
| 345 | ||
| 346 | fun find_rules goal rs = NetRules.retrieve rs (Logic.strip_assums_concl goal); | |
| 347 | ||
| 348 | fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => | |
| 349 | let | |
| 350 | fun get k = nth_elem (k, LocalRules.get ctxt); | |
| 351 | val rules = | |
| 352 | if not (null arg_rules) then arg_rules | |
| 353 | else find_erules facts (get elim_bangK) @ find_erules facts (get elimK) @ | |
| 354 | find_rules goal (get intro_bangK) @ find_rules goal (get introK); | |
| 355 | in trace rules; tac rules facts i end); | |
| 10309 | 356 | |
| 10744 | 357 | fun meth tac x = METHOD (HEADGOAL o tac x); | 
| 358 | fun meth' tac x y = METHOD (HEADGOAL o tac x y); | |
| 8220 | 359 | |
| 7419 | 360 | in | 
| 361 | ||
| 10744 | 362 | val rule_tac = gen_rule_tac Tactic.resolve_tac; | 
| 363 | val rule = meth rule_tac; | |
| 364 | val some_rule_tac = gen_some_rule_tac rule_tac; | |
| 365 | val some_rule = meth' some_rule_tac; | |
| 366 | ||
| 367 | val erule = meth' (gen_arule_tac Tactic.eresolve_tac); | |
| 368 | val drule = meth' (gen_arule_tac Tactic.dresolve_tac); | |
| 369 | val frule = meth' (gen_arule_tac Tactic.forward_tac); | |
| 5824 | 370 | |
| 7419 | 371 | end; | 
| 372 | ||
| 373 | ||
| 8195 | 374 | (* this *) | 
| 375 | ||
| 8671 | 376 | val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); | 
| 8195 | 377 | |
| 378 | ||
| 379 | (* assumption *) | |
| 7555 | 380 | |
| 10378 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 381 | fun asm_tac ths = | 
| 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 382 | foldr (op APPEND') (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths, K no_tac); | 
| 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 383 | |
| 10405 | 384 | fun assm_tac ctxt = | 
| 385 | assume_tac APPEND' | |
| 386 | asm_tac (ProofContext.prems_of ctxt) APPEND' | |
| 387 | Tactic.rtac Drule.reflexive_thm; | |
| 7419 | 388 | |
| 7555 | 389 | fun assumption_tac ctxt [] = assm_tac ctxt | 
| 10378 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 390 | | assumption_tac _ [fact] = asm_tac [fact] | 
| 7555 | 391 | | assumption_tac _ _ = K no_tac; | 
| 7419 | 392 | |
| 8671 | 393 | fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt); | 
| 7419 | 394 | |
| 395 | ||
| 9539 | 396 | (* res_inst_tac etc. *) | 
| 8238 | 397 | |
| 9539 | 398 | (*Note: insts refer to the implicit (!!) goal context; use at your own risk*) | 
| 9565 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 399 | fun gen_res_inst _ tac (quant, ([], thms)) = | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 400 | METHOD (fn facts => (quant (insert_tac facts THEN' tac thms))) | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 401 | | gen_res_inst tac _ (quant, (insts, [thm])) = | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 402 | METHOD (fn facts => (quant (insert_tac facts THEN' tac insts thm))) | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 403 | | gen_res_inst _ _ _ = error "Cannot have instantiations with multiple rules"; | 
| 8238 | 404 | |
| 9565 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 405 | val res_inst = gen_res_inst Tactic.res_inst_tac Tactic.resolve_tac; | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 406 | val eres_inst = gen_res_inst Tactic.eres_inst_tac Tactic.eresolve_tac; | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 407 | val dres_inst = gen_res_inst Tactic.dres_inst_tac Tactic.dresolve_tac; | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 408 | val forw_inst = gen_res_inst Tactic.forw_inst_tac Tactic.forward_tac; | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 409 | val cut_inst = gen_res_inst Tactic.cut_inst_tac Tactic.cut_facts_tac; | 
| 8238 | 410 | |
| 411 | ||
| 8329 | 412 | (* simple Prolog interpreter *) | 
| 413 | ||
| 414 | fun prolog_tac rules facts = | |
| 415 | DEPTH_SOLVE_1 (HEADGOAL (Tactic.assume_tac APPEND' Tactic.resolve_tac (facts @ rules))); | |
| 416 | ||
| 417 | val prolog = METHOD o prolog_tac; | |
| 418 | ||
| 419 | ||
| 8351 | 420 | (* ML tactics *) | 
| 421 | ||
| 422 | val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic); | |
| 423 | fun set_tactic f = tactic_ref := f; | |
| 424 | ||
| 425 | fun tactic txt ctxt = METHOD (fn facts => | |
| 9631 | 426 | (Context.use_mltext | 
| 427 |     ("let fun tactic (ctxt: PureIsar.Proof.context) (facts: thm list) : tactic = \
 | |
| 428 | \let val thm = PureIsar.ProofContext.get_thm_closure ctxt\n\ | |
| 429 | \ and thms = PureIsar.ProofContext.get_thms_closure ctxt in\n" | |
| 430 | ^ txt ^ | |
| 431 | "\nend in PureIsar.Method.set_tactic tactic end") | |
| 432 | false None; | |
| 433 | Context.setmp (Some (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts)); | |
| 8351 | 434 | |
| 435 | ||
| 5824 | 436 | |
| 437 | (** methods theory data **) | |
| 438 | ||
| 439 | (* data kind 'Isar/methods' *) | |
| 440 | ||
| 441 | structure MethodsDataArgs = | |
| 442 | struct | |
| 443 | val name = "Isar/methods"; | |
| 444 | type T = | |
| 445 |     {space: NameSpace.T,
 | |
| 446 | meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table}; | |
| 447 | ||
| 448 |   val empty = {space = NameSpace.empty, meths = Symtab.empty};
 | |
| 6546 | 449 | val copy = I; | 
| 5824 | 450 | val prep_ext = I; | 
| 451 |   fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) =
 | |
| 452 |     {space = NameSpace.merge (space1, space2),
 | |
| 453 | meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups => | |
| 454 |         error ("Attempt to merge different versions of methods " ^ commas_quote dups)};
 | |
| 455 | ||
| 9222 | 456 |   fun print _ {space, meths} =
 | 
| 5824 | 457 | let | 
| 458 | fun prt_meth (name, ((_, comment), _)) = Pretty.block | |
| 6849 | 459 | [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | 
| 5824 | 460 | in | 
| 8720 | 461 | [Pretty.big_list "methods:" (map prt_meth (NameSpace.cond_extern_table space meths))] | 
| 9222 | 462 | |> Pretty.chunks |> Pretty.writeln | 
| 5824 | 463 | end; | 
| 464 | end; | |
| 465 | ||
| 466 | structure MethodsData = TheoryDataFun(MethodsDataArgs); | |
| 467 | val print_methods = MethodsData.print; | |
| 7611 | 468 | |
| 5824 | 469 | |
| 470 | (* get methods *) | |
| 471 | ||
| 5916 | 472 | exception METHOD_FAIL of (string * Position.T) * exn; | 
| 473 | ||
| 5824 | 474 | fun method thy = | 
| 475 | let | |
| 476 |     val {space, meths} = MethodsData.get thy;
 | |
| 477 | ||
| 5884 | 478 | fun meth src = | 
| 479 | let | |
| 480 | val ((raw_name, _), pos) = Args.dest_src src; | |
| 481 | val name = NameSpace.intern space raw_name; | |
| 482 | in | |
| 5824 | 483 | (case Symtab.lookup (meths, name) of | 
| 484 |           None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
 | |
| 5916 | 485 | | Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src)) | 
| 5824 | 486 | end; | 
| 487 | in meth end; | |
| 488 | ||
| 489 | ||
| 9194 | 490 | (* add_method(s) *) | 
| 5824 | 491 | |
| 492 | fun add_methods raw_meths thy = | |
| 493 | let | |
| 494 | val full = Sign.full_name (Theory.sign_of thy); | |
| 495 | val new_meths = | |
| 496 | map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths; | |
| 497 | ||
| 498 |     val {space, meths} = MethodsData.get thy;
 | |
| 499 | val space' = NameSpace.extend (space, map fst new_meths); | |
| 500 | val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups => | |
| 501 |       error ("Duplicate declaration of method(s) " ^ commas_quote dups);
 | |
| 502 | in | |
| 503 |     thy |> MethodsData.put {space = space', meths = meths'}
 | |
| 504 | end; | |
| 505 | ||
| 9194 | 506 | val add_method = add_methods o Library.single; | 
| 507 | ||
| 5824 | 508 | (*implicit version*) | 
| 509 | fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]); | |
| 510 | ||
| 511 | ||
| 5884 | 512 | |
| 513 | (** method syntax **) | |
| 5824 | 514 | |
| 5884 | 515 | (* basic *) | 
| 516 | ||
| 517 | fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) = | |
| 518 | Args.syntax "method" scan; | |
| 5824 | 519 | |
| 8351 | 520 | fun simple_args scan f src ctxt : Proof.method = | 
| 521 | #2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); | |
| 522 | ||
| 7555 | 523 | fun ctxt_args (f: Proof.context -> Proof.method) src ctxt = | 
| 8282 | 524 | #2 (syntax (Scan.succeed (f ctxt)) src ctxt); | 
| 7555 | 525 | |
| 526 | fun no_args m = ctxt_args (K m); | |
| 5884 | 527 | |
| 528 | ||
| 529 | (* sections *) | |
| 5824 | 530 | |
| 7268 | 531 | type modifier = (Proof.context -> Proof.context) * Proof.context attribute; | 
| 532 | ||
| 533 | local | |
| 534 | ||
| 8381 | 535 | fun sect ss = Scan.first (map Scan.lift ss); | 
| 5884 | 536 | fun thms ss = Scan.unless (sect ss) Attrib.local_thms; | 
| 537 | fun thmss ss = Scan.repeat (thms ss) >> flat; | |
| 538 | ||
| 7268 | 539 | fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]); | 
| 5824 | 540 | |
| 7268 | 541 | fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt => | 
| 542 | Scan.succeed (apply m (ctxt, ths)))) >> #2; | |
| 5884 | 543 | |
| 7601 | 544 | fun sectioned args ss = args -- Scan.repeat (section ss); | 
| 5884 | 545 | |
| 7268 | 546 | in | 
| 5824 | 547 | |
| 5884 | 548 | fun sectioned_args args ss f src ctxt = | 
| 8282 | 549 | let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt | 
| 5921 | 550 | in f x ctxt' end; | 
| 5884 | 551 | |
| 7601 | 552 | fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; | 
| 9777 | 553 | fun bang_sectioned_args' ss scan f = | 
| 554 | sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); | |
| 7601 | 555 | fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); | 
| 7268 | 556 | |
| 8093 | 557 | fun thms_ctxt_args f = sectioned_args (thmss []) [] f; | 
| 558 | fun thms_args f = thms_ctxt_args (K o f); | |
| 9706 | 559 | fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); | 
| 5824 | 560 | |
| 7268 | 561 | end; | 
| 562 | ||
| 5824 | 563 | |
| 9539 | 564 | (* tactic syntax *) | 
| 8238 | 565 | |
| 10744 | 566 | fun nat_thms_args f = uncurry f oo | 
| 567 | (#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.local_thmss)); | |
| 568 | ||
| 8238 | 569 | val insts = | 
| 9539 | 570 | Scan.optional | 
| 9565 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 571 | (Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 572 | Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; | 
| 8238 | 573 | |
| 8537 | 574 | fun inst_args f = f oo (#2 oo syntax (Args.goal_spec HEADGOAL -- insts)); | 
| 575 | ||
| 576 | ||
| 9539 | 577 | fun goal_args' args tac = #2 oo syntax (Args.goal_spec HEADGOAL -- args >> | 
| 9706 | 578 | (fn (quant, s) => SIMPLE_METHOD' quant (tac s))); | 
| 8537 | 579 | |
| 9539 | 580 | fun goal_args args tac = goal_args' (Scan.lift args) tac; | 
| 8238 | 581 | |
| 582 | ||
| 5824 | 583 | |
| 584 | (** method text **) | |
| 585 | ||
| 586 | (* datatype text *) | |
| 587 | ||
| 588 | datatype text = | |
| 589 | Basic of (Proof.context -> Proof.method) | | |
| 590 | Source of Args.src | | |
| 591 | Then of text list | | |
| 592 | Orelse of text list | | |
| 593 | Try of text | | |
| 594 | Repeat1 of text; | |
| 595 | ||
| 596 | ||
| 597 | (* refine *) | |
| 598 | ||
| 8238 | 599 | fun gen_refine f text state = | 
| 5824 | 600 | let | 
| 601 | val thy = Proof.theory_of state; | |
| 602 | ||
| 8238 | 603 | fun eval (Basic mth) = f mth | 
| 604 | | eval (Source src) = f (method thy src) | |
| 5824 | 605 | | eval (Then txts) = Seq.EVERY (map eval txts) | 
| 606 | | eval (Orelse txts) = Seq.FIRST (map eval txts) | |
| 607 | | eval (Try txt) = Seq.TRY (eval txt) | |
| 608 | | eval (Repeat1 txt) = Seq.REPEAT1 (eval txt); | |
| 609 | in eval text state end; | |
| 610 | ||
| 8238 | 611 | val refine = gen_refine Proof.refine; | 
| 612 | val refine_end = gen_refine Proof.refine_end; | |
| 6404 | 613 | |
| 5824 | 614 | |
| 6404 | 615 | (* structured proof steps *) | 
| 5824 | 616 | |
| 7506 | 617 | val default_text = Source (Args.src (("default", []), Position.none));
 | 
| 8195 | 618 | val this_text = Basic (K this); | 
| 9706 | 619 | val done_text = Basic (K (SIMPLE_METHOD all_tac)); | 
| 7555 | 620 | |
| 8966 | 621 | fun close_text asm = Basic (fn ctxt => METHOD (K | 
| 622 | (FILTER Thm.no_prems ((if asm then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)))); | |
| 623 | ||
| 624 | fun finish_text asm None = close_text asm | |
| 625 | | finish_text asm (Some txt) = Then [txt, close_text asm]; | |
| 6872 | 626 | |
| 5824 | 627 | fun proof opt_text state = | 
| 628 | state | |
| 629 | |> Proof.assert_backward | |
| 6404 | 630 | |> refine (if_none opt_text default_text) | 
| 8242 | 631 | |> Seq.map (Proof.goal_facts (K [])) | 
| 5824 | 632 | |> Seq.map Proof.enter_forward; | 
| 633 | ||
| 8966 | 634 | fun local_qed asm opt_text = Proof.local_qed (refine (finish_text asm opt_text)); | 
| 635 | fun local_terminal_proof (text, opt_text) pr = | |
| 636 | Seq.THEN (proof (Some text), local_qed true opt_text pr); | |
| 637 | val local_default_proof = local_terminal_proof (default_text, None); | |
| 8195 | 638 | val local_immediate_proof = local_terminal_proof (this_text, None); | 
| 8966 | 639 | fun local_done_proof pr = Seq.THEN (proof (Some done_text), local_qed false None pr); | 
| 5824 | 640 | |
| 6872 | 641 | |
| 8966 | 642 | fun global_qeds asm opt_text = Proof.global_qed (refine (finish_text asm opt_text)); | 
| 5824 | 643 | |
| 8966 | 644 | fun global_qed asm opt_text state = | 
| 6872 | 645 | state | 
| 8966 | 646 | |> global_qeds asm opt_text | 
| 6872 | 647 | |> Proof.check_result "Failed to finish proof" state | 
| 648 | |> Seq.hd; | |
| 649 | ||
| 8966 | 650 | fun global_term_proof asm (text, opt_text) state = | 
| 6872 | 651 | state | 
| 652 | |> proof (Some text) | |
| 653 | |> Proof.check_result "Terminal proof method failed" state | |
| 8966 | 654 | |> (Seq.flat o Seq.map (global_qeds asm opt_text)) | 
| 6872 | 655 | |> Proof.check_result "Failed to finish proof (after successful terminal method)" state | 
| 656 | |> Seq.hd; | |
| 657 | ||
| 8966 | 658 | val global_terminal_proof = global_term_proof true; | 
| 6934 | 659 | val global_default_proof = global_terminal_proof (default_text, None); | 
| 8966 | 660 | val global_immediate_proof = global_terminal_proof (this_text, None); | 
| 661 | val global_done_proof = global_term_proof false (done_text, None); | |
| 5824 | 662 | |
| 663 | ||
| 664 | (** theory setup **) | |
| 665 | ||
| 9539 | 666 | (* misc tactic emulations *) | 
| 667 | ||
| 668 | val subgoal_meth = goal_args (Scan.repeat1 Args.name) Tactic.subgoals_tac; | |
| 669 | val thin_meth = goal_args Args.name Tactic.thin_tac; | |
| 670 | val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac; | |
| 9631 | 671 | val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac; | 
| 9539 | 672 | |
| 673 | ||
| 5824 | 674 | (* pure_methods *) | 
| 675 | ||
| 676 | val pure_methods = | |
| 677 |  [("fail", no_args fail, "force failure"),
 | |
| 678 |   ("succeed", no_args succeed, "succeed"),
 | |
| 9587 | 679 |   ("-", no_args insert_facts, "do nothing (insert current facts only)"),
 | 
| 9539 | 680 |   ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
 | 
| 7601 | 681 |   ("unfold", thms_args unfold, "unfold definitions"),
 | 
| 682 |   ("fold", thms_args fold, "fold definitions"),
 | |
| 12007 | 683 |   ("atomize", no_args (SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o ObjectLogic.atomize_tac)),
 | 
| 11962 | 684 | "present local premises as object-level statements"), | 
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 685 |   ("rule", thms_ctxt_args some_rule, "apply some rule"),
 | 
| 10744 | 686 |   ("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"),
 | 
| 687 |   ("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"),
 | |
| 688 |   ("frule", nat_thms_args frule, "apply rule in forward manner (improper)"),
 | |
| 8195 | 689 |   ("this", no_args this, "apply current facts as rules"),
 | 
| 8238 | 690 |   ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
 | 
| 9539 | 691 |   ("rule_tac", inst_args res_inst, "apply rule (dynamic instantiation!)"),
 | 
| 692 |   ("erule_tac", inst_args eres_inst, "apply rule in elimination manner (dynamic instantiation!)"),
 | |
| 693 |   ("drule_tac", inst_args dres_inst, "apply rule in destruct manner (dynamic instantiation!)"),
 | |
| 694 |   ("frule_tac", inst_args forw_inst, "apply rule in forward manner (dynamic instantiation!)"),
 | |
| 695 |   ("cut_tac", inst_args cut_inst, "cut rule (dynamic instantiation!)"),
 | |
| 9565 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 696 |   ("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation!)"),
 | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 697 |   ("thin_tac", thin_meth, "remove premise (dynamic instantiation!)"),
 | 
| 9631 | 698 |   ("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation!)"),
 | 
| 699 |   ("rotate_tac", rotate_meth, "rotate assumptions of goal"),
 | |
| 8351 | 700 |   ("prolog", thms_args prolog, "simple prolog interpreter"),
 | 
| 701 |   ("tactic", simple_args Args.name tactic, "ML tactic as proof method")];
 | |
| 5824 | 702 | |
| 703 | ||
| 704 | (* setup *) | |
| 705 | ||
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 706 | val setup = | 
| 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 707 | [GlobalRules.init, LocalRules.init, Attrib.add_attributes rule_atts, | 
| 11971 | 708 | MethodsData.init, add_methods pure_methods, | 
| 709 |   (#1 o PureThy.add_thms [(("", Drule.equal_intr_rule), [intro_global])])];
 | |
| 5824 | 710 | |
| 711 | ||
| 712 | end; | |
| 713 | ||
| 714 | ||
| 715 | structure BasicMethod: BASIC_METHOD = Method; | |
| 716 | open BasicMethod; |