| 5824 |      1 | (*  Title:      Pure/Isar/method.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Markus Wenzel, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | Proof methods.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | signature BASIC_METHOD =
 | 
|  |      9 | sig
 | 
|  |     10 |   val print_methods: theory -> unit
 | 
|  |     11 |   val Method: bstring -> (Args.src -> Proof.context -> Proof.method) -> string -> unit
 | 
|  |     12 | end;
 | 
|  |     13 | 
 | 
|  |     14 | signature METHOD =
 | 
|  |     15 | sig
 | 
|  |     16 |   include BASIC_METHOD
 | 
| 6500 |     17 |   val FINISHED: tactic -> tactic
 | 
| 6091 |     18 |   val LIFT: tactic -> thm -> (thm * (indexname * term) list * (string * thm list) list) Seq.seq
 | 
|  |     19 |   val METHOD: (thm list -> tactic) -> Proof.method
 | 
| 5824 |     20 |   val METHOD0: tactic -> Proof.method
 | 
|  |     21 |   val fail: Proof.method
 | 
|  |     22 |   val succeed: Proof.method
 | 
| 6500 |     23 |   val same_tac: thm list -> int -> tactic
 | 
| 6108 |     24 |   val same: Proof.method
 | 
| 5824 |     25 |   val assumption: Proof.method
 | 
|  |     26 |   val forward_chain: thm list -> thm list -> thm Seq.seq
 | 
| 6091 |     27 |   val rule_tac: thm list -> thm list -> int -> tactic
 | 
|  |     28 |   val rule: thm list -> Proof.method
 | 
| 5916 |     29 |   exception METHOD_FAIL of (string * Position.T) * exn
 | 
| 5824 |     30 |   val method: theory -> Args.src -> Proof.context -> Proof.method
 | 
|  |     31 |   val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list
 | 
|  |     32 |     -> theory -> theory
 | 
| 5884 |     33 |   val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
 | 
|  |     34 |     Proof.context -> Args.src -> Proof.context * 'a
 | 
|  |     35 |   val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method
 | 
|  |     36 |   val sectioned_args: ((Args.T list -> Proof.context attribute * Args.T list) list ->
 | 
|  |     37 |       Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
 | 
|  |     38 |     (Args.T list -> Proof.context attribute * Args.T list) list ->
 | 
|  |     39 |     ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
 | 
|  |     40 |   val default_sectioned_args: Proof.context attribute ->
 | 
|  |     41 |     (Args.T list -> Proof.context attribute * Args.T list) list ->
 | 
|  |     42 |     (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
 | 
|  |     43 |   val only_sectioned_args: (Args.T list -> Proof.context attribute * Args.T list) list ->
 | 
|  |     44 |     (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
 | 
| 6091 |     45 |   val thms_args: (thm list -> Proof.method) -> Args.src -> Proof.context -> Proof.method
 | 
| 5824 |     46 |   datatype text =
 | 
|  |     47 |     Basic of (Proof.context -> Proof.method) |
 | 
|  |     48 |     Source of Args.src |
 | 
|  |     49 |     Then of text list |
 | 
|  |     50 |     Orelse of text list |
 | 
|  |     51 |     Try of text |
 | 
|  |     52 |     Repeat of text |
 | 
|  |     53 |     Repeat1 of text
 | 
|  |     54 |   val refine: text -> Proof.state -> Proof.state Seq.seq
 | 
|  |     55 |   val tac: text -> Proof.state -> Proof.state Seq.seq
 | 
| 5884 |     56 |   val then_tac: text -> Proof.state -> Proof.state Seq.seq
 | 
| 5824 |     57 |   val proof: text option -> Proof.state -> Proof.state Seq.seq
 | 
| 6404 |     58 |   val local_qed: text option -> Proof.state -> Proof.state Seq.seq
 | 
|  |     59 |   val local_terminal_proof: text -> Proof.state -> Proof.state Seq.seq
 | 
|  |     60 |   val local_immediate_proof: Proof.state -> Proof.state Seq.seq
 | 
|  |     61 |   val local_default_proof: Proof.state -> Proof.state Seq.seq
 | 
|  |     62 |   val global_qed: bstring option -> theory attribute list option -> text option
 | 
| 6532 |     63 |     -> Proof.state -> theory * {kind: string, name: string, thm: thm}
 | 
|  |     64 |   val global_terminal_proof: text -> Proof.state -> theory * {kind: string, name: string, thm: thm}
 | 
|  |     65 |   val global_immediate_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
 | 
|  |     66 |   val global_default_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
 | 
| 5824 |     67 |   val setup: (theory -> theory) list
 | 
|  |     68 | end;
 | 
|  |     69 | 
 | 
|  |     70 | structure Method: METHOD =
 | 
|  |     71 | struct
 | 
|  |     72 | 
 | 
|  |     73 | 
 | 
| 6500 |     74 | (** utils **)
 | 
|  |     75 | 
 | 
|  |     76 | val FINISHED = FILTER (equal 0 o Thm.nprems_of);
 | 
|  |     77 | 
 | 
|  |     78 | 
 | 
|  |     79 | 
 | 
| 5824 |     80 | (** proof methods **)
 | 
|  |     81 | 
 | 
|  |     82 | (* method from tactic *)
 | 
|  |     83 | 
 | 
|  |     84 | fun LIFT tac goal = Seq.map (fn x => (x, [], [])) (tac goal);
 | 
|  |     85 | fun METHOD tacf = Proof.method (LIFT o tacf);
 | 
|  |     86 | fun METHOD0 tac = METHOD (fn [] => tac | _ => error "Method may not be used with facts");
 | 
|  |     87 | 
 | 
|  |     88 | 
 | 
|  |     89 | (* primitive *)
 | 
|  |     90 | 
 | 
|  |     91 | val fail = METHOD (K no_tac);
 | 
|  |     92 | val succeed = METHOD (K all_tac);
 | 
|  |     93 | 
 | 
|  |     94 | 
 | 
| 6108 |     95 | (* same, assumption *)
 | 
| 5824 |     96 | 
 | 
| 6108 |     97 | fun same_tac [] = K all_tac
 | 
|  |     98 |   | same_tac facts =
 | 
| 6091 |     99 |       let val r = ~ (length facts);
 | 
|  |    100 |       in metacuts_tac facts THEN' rotate_tac r end;
 | 
| 5824 |    101 | 
 | 
| 6108 |    102 | val same = METHOD (ALLGOALS o same_tac);
 | 
|  |    103 | val assumption = METHOD (fn facts => FIRSTGOAL (same_tac facts THEN' assume_tac));
 | 
| 5824 |    104 | 
 | 
| 6500 |    105 | val asm_finish = METHOD (K (FINISHED (ALLGOALS assume_tac)));
 | 
| 5824 |    106 | 
 | 
|  |    107 | 
 | 
| 6532 |    108 | (* fold / unfold definitions *)
 | 
|  |    109 | 
 | 
|  |    110 | val fold = METHOD0 o fold_goals_tac;
 | 
|  |    111 | val unfold = METHOD0 o rewrite_goals_tac;
 | 
|  |    112 | 
 | 
|  |    113 | 
 | 
| 5824 |    114 | (* rule *)
 | 
|  |    115 | 
 | 
|  |    116 | fun res th i rule = Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty;
 | 
|  |    117 | 
 | 
|  |    118 | fun multi_resolve facts rule =
 | 
|  |    119 |   let
 | 
|  |    120 |     fun multi_res i [] = Seq.single rule
 | 
|  |    121 |       | multi_res i (th :: ths) = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths))
 | 
|  |    122 |   in multi_res 1 facts end;
 | 
|  |    123 | 
 | 
|  |    124 | fun forward_chain facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules));
 | 
|  |    125 | 
 | 
| 6091 |    126 | fun rule_tac rules [] = resolve_tac rules
 | 
| 5824 |    127 |   | rule_tac erules facts =
 | 
|  |    128 |       let
 | 
| 6091 |    129 |         val rules = forward_chain facts erules;
 | 
| 5824 |    130 |         fun tac i state = Seq.flat (Seq.map (fn rule => Tactic.rtac rule i state) rules);
 | 
|  |    131 |       in tac end;
 | 
|  |    132 | 
 | 
|  |    133 | fun rule rules = METHOD (FIRSTGOAL o rule_tac rules);
 | 
|  |    134 | 
 | 
|  |    135 | 
 | 
|  |    136 | 
 | 
|  |    137 | (** methods theory data **)
 | 
|  |    138 | 
 | 
|  |    139 | (* data kind 'Isar/methods' *)
 | 
|  |    140 | 
 | 
|  |    141 | structure MethodsDataArgs =
 | 
|  |    142 | struct
 | 
|  |    143 |   val name = "Isar/methods";
 | 
|  |    144 |   type T =
 | 
|  |    145 |     {space: NameSpace.T,
 | 
|  |    146 |      meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table};
 | 
|  |    147 | 
 | 
|  |    148 |   val empty = {space = NameSpace.empty, meths = Symtab.empty};
 | 
| 6546 |    149 |   val copy = I;
 | 
| 5824 |    150 |   val prep_ext = I;
 | 
|  |    151 |   fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) =
 | 
|  |    152 |     {space = NameSpace.merge (space1, space2),
 | 
|  |    153 |       meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups =>
 | 
|  |    154 |         error ("Attempt to merge different versions of methods " ^ commas_quote dups)};
 | 
|  |    155 | 
 | 
|  |    156 |   fun print _ {space, meths} =
 | 
|  |    157 |     let
 | 
|  |    158 |       fun prt_meth (name, ((_, comment), _)) = Pretty.block
 | 
|  |    159 |         [Pretty.str (NameSpace.cond_extern space name ^ ":"), Pretty.brk 2, Pretty.str comment];
 | 
|  |    160 |     in
 | 
|  |    161 |       Pretty.writeln (Display.pretty_name_space ("method name space", space));
 | 
|  |    162 |       Pretty.writeln (Pretty.big_list "methods:" (map prt_meth (Symtab.dest meths)))
 | 
|  |    163 |     end;
 | 
|  |    164 | end;
 | 
|  |    165 | 
 | 
|  |    166 | structure MethodsData = TheoryDataFun(MethodsDataArgs);
 | 
|  |    167 | val print_methods = MethodsData.print;
 | 
|  |    168 | 
 | 
|  |    169 | 
 | 
|  |    170 | (* get methods *)
 | 
|  |    171 | 
 | 
| 5916 |    172 | exception METHOD_FAIL of (string * Position.T) * exn;
 | 
|  |    173 | 
 | 
| 5824 |    174 | fun method thy =
 | 
|  |    175 |   let
 | 
|  |    176 |     val {space, meths} = MethodsData.get thy;
 | 
|  |    177 | 
 | 
| 5884 |    178 |     fun meth src =
 | 
|  |    179 |       let
 | 
|  |    180 |         val ((raw_name, _), pos) = Args.dest_src src;
 | 
|  |    181 |         val name = NameSpace.intern space raw_name;
 | 
|  |    182 |       in
 | 
| 5824 |    183 |         (case Symtab.lookup (meths, name) of
 | 
|  |    184 |           None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
 | 
| 5916 |    185 |         | Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src))
 | 
| 5824 |    186 |       end;
 | 
|  |    187 |   in meth end;
 | 
|  |    188 | 
 | 
|  |    189 | 
 | 
|  |    190 | (* add_methods *)
 | 
|  |    191 | 
 | 
|  |    192 | fun add_methods raw_meths thy =
 | 
|  |    193 |   let
 | 
|  |    194 |     val full = Sign.full_name (Theory.sign_of thy);
 | 
|  |    195 |     val new_meths =
 | 
|  |    196 |       map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths;
 | 
|  |    197 | 
 | 
|  |    198 |     val {space, meths} = MethodsData.get thy;
 | 
|  |    199 |     val space' = NameSpace.extend (space, map fst new_meths);
 | 
|  |    200 |     val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups =>
 | 
|  |    201 |       error ("Duplicate declaration of method(s) " ^ commas_quote dups);
 | 
|  |    202 |   in
 | 
|  |    203 |     thy |> MethodsData.put {space = space', meths = meths'}
 | 
|  |    204 |   end;
 | 
|  |    205 | 
 | 
|  |    206 | (*implicit version*)
 | 
|  |    207 | fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]);
 | 
|  |    208 | 
 | 
|  |    209 | 
 | 
| 5884 |    210 | 
 | 
|  |    211 | (** method syntax **)
 | 
| 5824 |    212 | 
 | 
| 5884 |    213 | (* basic *)
 | 
|  |    214 | 
 | 
|  |    215 | fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) =
 | 
|  |    216 |   Args.syntax "method" scan;
 | 
| 5824 |    217 | 
 | 
| 5884 |    218 | fun no_args (x: Proof.method) src ctxt = #2 (syntax (Scan.succeed x) ctxt src);
 | 
|  |    219 | 
 | 
|  |    220 | 
 | 
|  |    221 | (* sections *)
 | 
| 5824 |    222 | 
 | 
| 5884 |    223 | fun sect ss = Scan.first (map (fn s => Scan.lift (s --| Args.$$$ ":")) ss);
 | 
|  |    224 | fun thms ss = Scan.unless (sect ss) Attrib.local_thms;
 | 
|  |    225 | fun thmss ss = Scan.repeat (thms ss) >> flat;
 | 
|  |    226 | 
 | 
| 6091 |    227 | fun apply att (ctxt, ths) = Thm.applys_attributes ((ctxt, ths), [att]);
 | 
| 5824 |    228 | 
 | 
| 5884 |    229 | fun section ss = (sect ss -- thmss ss) :-- (fn (att, ths) => Scan.depend (fn ctxt =>
 | 
|  |    230 |   Scan.succeed (apply att (ctxt, ths)))) >> #2;
 | 
|  |    231 | 
 | 
|  |    232 | fun sectioned args ss = args ss -- Scan.repeat (section ss);
 | 
|  |    233 | 
 | 
| 5824 |    234 | 
 | 
| 5884 |    235 | fun sectioned_args args ss f src ctxt =
 | 
| 5921 |    236 |   let val (ctxt', (x, _)) = syntax (sectioned args ss) ctxt src
 | 
|  |    237 |   in f x ctxt' end;
 | 
| 5884 |    238 | 
 | 
|  |    239 | fun default_sectioned_args att ss f src ctxt =
 | 
|  |    240 |   sectioned_args thmss ss (fn ths => fn ctxt' => f (#1 (apply att (ctxt', ths)))) src ctxt;
 | 
|  |    241 | 
 | 
|  |    242 | fun only_sectioned_args ss f = sectioned_args (K (Scan.succeed ())) ss (fn () => f);
 | 
|  |    243 | fun thms_args f = sectioned_args thmss [] (fn ths => fn _ => f ths);
 | 
| 5824 |    244 | 
 | 
|  |    245 | 
 | 
|  |    246 | 
 | 
|  |    247 | (** method text **)
 | 
|  |    248 | 
 | 
|  |    249 | (* datatype text *)
 | 
|  |    250 | 
 | 
|  |    251 | datatype text =
 | 
|  |    252 |   Basic of (Proof.context -> Proof.method) |
 | 
|  |    253 |   Source of Args.src |
 | 
|  |    254 |   Then of text list |
 | 
|  |    255 |   Orelse of text list |
 | 
|  |    256 |   Try of text |
 | 
|  |    257 |   Repeat of text |
 | 
|  |    258 |   Repeat1 of text;
 | 
|  |    259 | 
 | 
|  |    260 | 
 | 
|  |    261 | (* refine *)
 | 
|  |    262 | 
 | 
|  |    263 | fun refine text state =
 | 
|  |    264 |   let
 | 
|  |    265 |     val thy = Proof.theory_of state;
 | 
|  |    266 | 
 | 
|  |    267 |     fun eval (Basic mth) = Proof.refine mth
 | 
|  |    268 |       | eval (Source src) = Proof.refine (method thy src)
 | 
|  |    269 |       | eval (Then txts) = Seq.EVERY (map eval txts)
 | 
|  |    270 |       | eval (Orelse txts) = Seq.FIRST (map eval txts)
 | 
|  |    271 |       | eval (Try txt) = Seq.TRY (eval txt)
 | 
|  |    272 |       | eval (Repeat txt) = Seq.REPEAT (eval txt)
 | 
|  |    273 |       | eval (Repeat1 txt) = Seq.REPEAT1 (eval txt);
 | 
|  |    274 |   in eval text state end;
 | 
|  |    275 | 
 | 
| 6404 |    276 | fun named_method name = Source (Args.src ((name, []), Position.none));
 | 
|  |    277 | 
 | 
| 5824 |    278 | 
 | 
|  |    279 | (* unstructured steps *)
 | 
|  |    280 | 
 | 
|  |    281 | fun tac text state =
 | 
|  |    282 |   state
 | 
|  |    283 |   |> Proof.goal_facts (K [])
 | 
|  |    284 |   |> refine text;
 | 
|  |    285 | 
 | 
| 5884 |    286 | fun then_tac text state =
 | 
| 5824 |    287 |   state
 | 
|  |    288 |   |> Proof.goal_facts Proof.the_facts
 | 
|  |    289 |   |> refine text;
 | 
|  |    290 | 
 | 
|  |    291 | 
 | 
| 6404 |    292 | (* structured proof steps *)
 | 
| 5824 |    293 | 
 | 
| 6404 |    294 | val default_text = named_method "default";
 | 
|  |    295 | val finish_text = named_method "finish";
 | 
| 5824 |    296 | 
 | 
|  |    297 | fun proof opt_text state =
 | 
|  |    298 |   state
 | 
|  |    299 |   |> Proof.assert_backward
 | 
| 6404 |    300 |   |> refine (if_none opt_text default_text)
 | 
| 5824 |    301 |   |> Seq.map Proof.enter_forward;
 | 
|  |    302 | 
 | 
| 6404 |    303 | fun local_qed opt_text = Proof.local_qed (refine (if_none opt_text finish_text));
 | 
|  |    304 | fun local_terminal_proof text = Seq.THEN (proof (Some text), local_qed None);
 | 
|  |    305 | val local_immediate_proof = local_terminal_proof (Basic (K same));
 | 
|  |    306 | val local_default_proof = local_terminal_proof default_text;
 | 
| 5824 |    307 | 
 | 
| 6404 |    308 | fun global_qed alt_name alt_atts opt_text =
 | 
|  |    309 |   Proof.global_qed (refine (if_none opt_text finish_text)) alt_name alt_atts;
 | 
| 5824 |    310 | 
 | 
| 6404 |    311 | fun global_terminal_proof text = Seq.hd o Seq.map (global_qed None None None) o proof (Some text);
 | 
|  |    312 | val global_immediate_proof = global_terminal_proof (Basic (K same));
 | 
|  |    313 | val global_default_proof = global_terminal_proof default_text;
 | 
| 5824 |    314 | 
 | 
|  |    315 | 
 | 
|  |    316 | 
 | 
|  |    317 | (** theory setup **)
 | 
|  |    318 | 
 | 
|  |    319 | (* pure_methods *)
 | 
|  |    320 | 
 | 
|  |    321 | val pure_methods =
 | 
|  |    322 |  [("fail", no_args fail, "force failure"),
 | 
|  |    323 |   ("succeed", no_args succeed, "succeed"),
 | 
| 6108 |    324 |   ("same", no_args same, "insert facts, nothing else"),
 | 
| 5824 |    325 |   ("assumption", no_args assumption, "proof by assumption"),
 | 
|  |    326 |   ("finish", no_args asm_finish, "finish proof by assumption"),
 | 
| 6532 |    327 |   ("fold", thms_args fold, "fold definitions"),
 | 
|  |    328 |   ("unfold", thms_args unfold, "unfold definitions"),
 | 
| 5884 |    329 |   ("rule", thms_args rule, "apply primitive rule")];
 | 
| 5824 |    330 | 
 | 
|  |    331 | 
 | 
|  |    332 | (* setup *)
 | 
|  |    333 | 
 | 
|  |    334 | val setup = [MethodsData.init, add_methods pure_methods];
 | 
|  |    335 | 
 | 
|  |    336 | 
 | 
|  |    337 | end;
 | 
|  |    338 | 
 | 
|  |    339 | 
 | 
|  |    340 | structure BasicMethod: BASIC_METHOD = Method;
 | 
|  |    341 | open BasicMethod;
 |