src/Pure/Isar/method.ML
author wenzelm
Thu, 02 Mar 2000 18:18:31 +0100
changeset 8329 8308b7a152a3
parent 8282 58a33fd5b30c
child 8335 4c117393e4e6
permissions -rw-r--r--
added 'prolog' method;
     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
    17   val print_global_rules: theory -> unit
    18   val print_local_rules: Proof.context -> unit
    19   val dest_global: theory attribute
    20   val elim_global: theory attribute
    21   val intro_global: theory attribute
    22   val delrule_global: theory attribute
    23   val dest_local: Proof.context attribute
    24   val elim_local: Proof.context attribute
    25   val intro_local: Proof.context attribute
    26   val delrule_local: Proof.context attribute
    27   val METHOD: (thm list -> tactic) -> Proof.method
    28   val METHOD0: tactic -> Proof.method
    29   val fail: Proof.method
    30   val succeed: Proof.method
    31   val defer: int option -> Proof.method
    32   val prefer: int -> Proof.method
    33   val insert_tac: thm list -> int -> tactic
    34   val insert: thm list -> Proof.method
    35   val insert_facts: Proof.method
    36   val unfold: thm list -> Proof.method
    37   val fold: thm list -> Proof.method
    38   val multi_resolve: thm list -> thm -> thm Seq.seq
    39   val multi_resolves: thm list -> thm list -> thm Seq.seq
    40   val rule_tac: thm list -> thm list -> int -> tactic
    41   val rule: thm list -> Proof.method
    42   val erule: thm list -> Proof.method
    43   val drule: thm list -> Proof.method
    44   val frule: thm list -> Proof.method
    45   val this: Proof.method
    46   val assumption: Proof.context -> Proof.method
    47   exception METHOD_FAIL of (string * Position.T) * exn
    48   val help_methods: theory option -> unit
    49   val method: theory -> Args.src -> Proof.context -> Proof.method
    50   val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list
    51     -> theory -> theory
    52   val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
    53     Args.src -> Proof.context -> Proof.context * 'a
    54   val ctxt_args: (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
    55   val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method
    56   type modifier
    57   val sectioned_args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
    58     (Args.T list -> modifier * Args.T list) list ->
    59     ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
    60   val bang_sectioned_args:
    61     (Args.T list -> modifier * Args.T list) list ->
    62     (thm list -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
    63   val only_sectioned_args:
    64     (Args.T list -> modifier * Args.T list) list ->
    65     (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
    66   val thms_ctxt_args: (thm list -> Proof.context -> Proof.method)
    67     -> Args.src -> Proof.context -> Proof.method
    68   val thms_args: (thm list -> Proof.method) -> Args.src -> Proof.context -> Proof.method
    69   datatype text =
    70     Basic of (Proof.context -> Proof.method) |
    71     Source of Args.src |
    72     Then of text list |
    73     Orelse of text list |
    74     Try of text |
    75     Repeat1 of text
    76   val refine: text -> Proof.state -> Proof.state Seq.seq
    77   val refine_end: text -> Proof.state -> Proof.state Seq.seq
    78   val proof: text option -> Proof.state -> Proof.state Seq.seq
    79   val local_qed: text option
    80     -> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
    81     -> Proof.state -> Proof.state Seq.seq
    82   val local_terminal_proof: text * text option
    83     -> ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
    84     -> Proof.state -> Proof.state Seq.seq
    85   val local_immediate_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
    86     -> Proof.state -> Proof.state Seq.seq
    87   val local_default_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
    88     -> Proof.state -> Proof.state Seq.seq
    89   val global_qed: text option -> Proof.state -> theory * {kind: string, name: string, thm: thm}
    90   val global_terminal_proof: text * text option
    91     -> Proof.state -> theory * {kind: string, name: string, thm: thm}
    92   val global_immediate_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
    93   val global_default_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
    94   val setup: (theory -> theory) list
    95 end;
    96 
    97 structure Method: METHOD =
    98 struct
    99 
   100 
   101 (** global and local rule data **)
   102 
   103 fun prt_rules kind ths =
   104   Pretty.writeln (Pretty.big_list ("standard " ^ kind ^ " rules:") (map Display.pretty_thm ths));
   105 
   106 fun print_rules (intro, elim) =
   107   (prt_rules "introduction" intro; prt_rules "elimination" elim);
   108 
   109 fun merge_rules (ths1, ths2) = Library.generic_merge Thm.eq_thm I I ths1 ths2;
   110 
   111 
   112 (* theory data kind 'Isar/rules' *)
   113 
   114 structure GlobalRulesArgs =
   115 struct
   116   val name = "Isar/rules";
   117   type T = thm list * thm list;
   118 
   119   val empty = ([], []);
   120   val copy = I;
   121   val prep_ext = I;
   122   fun merge ((intro1, elim1), (intro2, elim2)) =
   123     (merge_rules (intro1, intro2), merge_rules (elim1, elim2));
   124   fun print _ = print_rules;
   125 end;
   126 
   127 structure GlobalRules = TheoryDataFun(GlobalRulesArgs);
   128 val print_global_rules = GlobalRules.print;
   129 
   130 
   131 (* proof data kind 'Isar/rules' *)
   132 
   133 structure LocalRulesArgs =
   134 struct
   135   val name = "Isar/rules";
   136   type T = thm list * thm list;
   137 
   138   val init = GlobalRules.get;
   139   fun print _ = print_rules;
   140 end;
   141 
   142 structure LocalRules = ProofDataFun(LocalRulesArgs);
   143 val print_local_rules = LocalRules.print;
   144 
   145 
   146 
   147 (** attributes **)
   148 
   149 (* add rules *)
   150 
   151 local
   152 
   153 fun add_rule thm rules = Library.gen_ins Thm.eq_thm (thm, rules);
   154 fun del_rule thm rules = Library.gen_rem Thm.eq_thm (rules, thm);
   155 
   156 fun add_dest thm (intro, elim) = (intro, add_rule (Tactic.make_elim thm) elim);
   157 fun add_elim thm (intro, elim) = (intro, add_rule thm elim);
   158 fun add_intro thm (intro, elim) = (add_rule thm intro, elim);
   159 fun delrule thm (intro, elim) = (del_rule thm intro, del_rule thm elim);
   160 
   161 fun mk_att f g (x, thm) = (f (g thm) x, thm);
   162 
   163 in
   164 
   165 val dest_global = mk_att GlobalRules.map add_dest;
   166 val elim_global = mk_att GlobalRules.map add_elim;
   167 val intro_global = mk_att GlobalRules.map add_intro;
   168 val delrule_global = mk_att GlobalRules.map delrule;
   169 
   170 val dest_local = mk_att LocalRules.map add_dest;
   171 val elim_local = mk_att LocalRules.map add_elim;
   172 val intro_local = mk_att LocalRules.map add_intro;
   173 val delrule_local = mk_att LocalRules.map delrule;
   174 
   175 end;
   176 
   177 
   178 (* concrete syntax *)
   179 
   180 val rule_atts =
   181  [("dest", (Attrib.no_args dest_global, Attrib.no_args dest_local), "destruction rule"),
   182   ("elim", (Attrib.no_args elim_global, Attrib.no_args elim_local), "elimination rule"),
   183   ("intro", (Attrib.no_args intro_global, Attrib.no_args intro_local), "introduction rule"),
   184   ("delrule", (Attrib.no_args delrule_global, Attrib.no_args delrule_local), "delete rule")];
   185 
   186 
   187 
   188 (** proof methods **)
   189 
   190 (* method from tactic *)
   191 
   192 val METHOD = Proof.method;
   193 fun METHOD0 tac = METHOD (fn [] => tac | _ => error "Method may not be used with facts");
   194 
   195 
   196 (* primitive *)
   197 
   198 val fail = METHOD (K no_tac);
   199 val succeed = METHOD (K all_tac);
   200 
   201 
   202 (* shuffle *)
   203 
   204 fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)));
   205 fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1)));
   206 
   207 
   208 (* insert *)
   209 
   210 local
   211 
   212 fun cut_rule_tac raw_rule =
   213   let
   214     val rule = Drule.forall_intr_vars raw_rule;
   215     val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl;
   216   in Tactic.rtac (rule COMP revcut_rl) end;
   217 
   218 in
   219 
   220 fun insert_tac [] i = all_tac
   221   | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts);
   222 
   223 val insert_facts = METHOD (ALLGOALS o insert_tac);
   224 fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms));
   225 
   226 end;
   227 
   228 
   229 (* unfold / fold definitions *)
   230 
   231 fun unfold thms = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN rewrite_goals_tac thms);
   232 fun fold thms = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN fold_goals_tac thms);
   233 
   234 
   235 (* multi_resolve *)
   236 
   237 local
   238 
   239 fun res th i rule =
   240   Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty;
   241 
   242 fun multi_res _ [] rule = Seq.single rule
   243   | multi_res i (th :: ths) rule = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths rule));
   244 
   245 in
   246 
   247 val multi_resolve = multi_res 1;
   248 fun multi_resolves facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules));
   249 
   250 end;
   251 
   252 
   253 (* rule *)
   254 
   255 local
   256 
   257 fun gen_rule_tac tac rules [] = tac rules
   258   | gen_rule_tac tac erules facts =
   259       let
   260         val rules = multi_resolves facts erules;
   261         fun tactic i state = Seq.flat (Seq.map (fn rule => tac [rule] i state) rules);
   262       in tactic end;
   263 
   264 fun gen_rule tac rules = METHOD (FINDGOAL o tac rules);
   265 
   266 fun gen_rule' tac arg_rules ctxt = METHOD (fn facts =>
   267   let val rules =
   268     if not (null arg_rules) then arg_rules
   269     else if null facts then #1 (LocalRules.get ctxt)
   270     else op @ (LocalRules.get ctxt);
   271   in FINDGOAL (tac rules facts) end);
   272 
   273 fun setup raw_tac =
   274   let val tac = gen_rule_tac raw_tac
   275   in (tac, gen_rule tac, gen_rule' tac) end;
   276 
   277 in
   278 
   279 val (rule_tac, rule, some_rule) = setup Tactic.resolve_tac;
   280 val (erule_tac, erule, some_erule) = setup Tactic.eresolve_tac;
   281 val (drule_tac, drule, some_drule) = setup Tactic.dresolve_tac;
   282 val (frule_tac, frule, some_frule) = setup Tactic.forward_tac;
   283 
   284 end;
   285 
   286 
   287 (* this *)
   288 
   289 val this = METHOD (EVERY o map (FINDGOAL o Tactic.rtac));
   290 
   291 
   292 (* assumption *)
   293 
   294 fun assm_tac ctxt =
   295   assume_tac APPEND' resolve_tac (filter Thm.no_prems (ProofContext.prems_of ctxt));
   296 
   297 fun assumption_tac ctxt [] = assm_tac ctxt
   298   | assumption_tac _ [fact] = resolve_tac [fact]
   299   | assumption_tac _ _ = K no_tac;
   300 
   301 fun assumption ctxt = METHOD (FINDGOAL o assumption_tac ctxt);
   302 
   303 
   304 (* res_inst_tac emulations *)
   305 
   306 (*Note: insts refer to the implicit (!) goal context; use at your own risk*)
   307 fun gen_res_inst tac ((i, insts), thm) =
   308   METHOD (fn facts => (insert_tac facts THEN' tac insts thm) i);
   309 
   310 val res_inst = gen_res_inst Tactic.res_inst_tac;
   311 val eres_inst = gen_res_inst Tactic.eres_inst_tac;
   312 val dres_inst = gen_res_inst Tactic.dres_inst_tac;
   313 val forw_inst = gen_res_inst Tactic.forw_inst_tac;
   314 
   315 
   316 (* simple Prolog interpreter *)
   317 
   318 fun prolog_tac rules facts =
   319   DEPTH_SOLVE_1 (HEADGOAL (Tactic.assume_tac APPEND' Tactic.resolve_tac (facts @ rules)));
   320 
   321 val prolog = METHOD o prolog_tac;
   322 
   323 
   324 
   325 (** methods theory data **)
   326 
   327 (* data kind 'Isar/methods' *)
   328 
   329 structure MethodsDataArgs =
   330 struct
   331   val name = "Isar/methods";
   332   type T =
   333     {space: NameSpace.T,
   334      meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table};
   335 
   336   val empty = {space = NameSpace.empty, meths = Symtab.empty};
   337   val copy = I;
   338   val prep_ext = I;
   339   fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) =
   340     {space = NameSpace.merge (space1, space2),
   341       meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups =>
   342         error ("Attempt to merge different versions of methods " ^ commas_quote dups)};
   343 
   344   fun print_meths verbose {space, meths} =
   345     let
   346       fun prt_meth (name, ((_, comment), _)) = Pretty.block
   347         [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
   348     in
   349       if not verbose then ()
   350       else Pretty.writeln (Display.pretty_name_space ("method name space", space));
   351       Pretty.writeln (Pretty.big_list "methods:"
   352         (map prt_meth (NameSpace.cond_extern_table space meths)))
   353     end;
   354 
   355   fun print _ = print_meths true;
   356 end;
   357 
   358 structure MethodsData = TheoryDataFun(MethodsDataArgs);
   359 val print_methods = MethodsData.print;
   360 
   361 fun help_methods None = writeln "methods: (unkown theory context)"
   362   | help_methods (Some thy) = MethodsDataArgs.print_meths false (MethodsData.get thy);
   363 
   364 
   365 (* get methods *)
   366 
   367 exception METHOD_FAIL of (string * Position.T) * exn;
   368 
   369 fun method thy =
   370   let
   371     val {space, meths} = MethodsData.get thy;
   372 
   373     fun meth src =
   374       let
   375         val ((raw_name, _), pos) = Args.dest_src src;
   376         val name = NameSpace.intern space raw_name;
   377       in
   378         (case Symtab.lookup (meths, name) of
   379           None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
   380         | Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src))
   381       end;
   382   in meth end;
   383 
   384 
   385 (* add_methods *)
   386 
   387 fun add_methods raw_meths thy =
   388   let
   389     val full = Sign.full_name (Theory.sign_of thy);
   390     val new_meths =
   391       map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths;
   392 
   393     val {space, meths} = MethodsData.get thy;
   394     val space' = NameSpace.extend (space, map fst new_meths);
   395     val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups =>
   396       error ("Duplicate declaration of method(s) " ^ commas_quote dups);
   397   in
   398     thy |> MethodsData.put {space = space', meths = meths'}
   399   end;
   400 
   401 (*implicit version*)
   402 fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]);
   403 
   404 
   405 
   406 (** method syntax **)
   407 
   408 (* basic *)
   409 
   410 fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) =
   411   Args.syntax "method" scan;
   412 
   413 fun ctxt_args (f: Proof.context -> Proof.method) src ctxt =
   414   #2 (syntax (Scan.succeed (f ctxt)) src ctxt);
   415 
   416 fun no_args m = ctxt_args (K m);
   417 
   418 
   419 (* sections *)
   420 
   421 type modifier = (Proof.context -> Proof.context) * Proof.context attribute;
   422 
   423 local
   424 
   425 fun sect ss = Scan.first (map (fn s => Scan.lift (s --| Args.$$$ ":")) ss);
   426 fun thms ss = Scan.unless (sect ss) Attrib.local_thms;
   427 fun thmss ss = Scan.repeat (thms ss) >> flat;
   428 
   429 fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]);
   430 
   431 fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt =>
   432   Scan.succeed (apply m (ctxt, ths)))) >> #2;
   433 
   434 fun sectioned args ss = args -- Scan.repeat (section ss);
   435 
   436 in
   437 
   438 fun sectioned_args args ss f src ctxt =
   439   let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt
   440   in f x ctxt' end;
   441 
   442 fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f;
   443 fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f);
   444 
   445 fun thms_ctxt_args f = sectioned_args (thmss []) [] f;
   446 fun thms_args f = thms_ctxt_args (K o f);
   447 
   448 end;
   449 
   450 
   451 (* insts *)
   452 
   453 val insts =
   454   Scan.lift (Scan.optional (Args.$$$ "(" |-- Args.!!! (Args.nat --| Args.$$$ ")")) 1) --
   455     Args.enum1 "and" (Scan.lift (Args.name -- Args.!!! (Args.$$$ "=" |-- Args.name))) --
   456     (Scan.lift (Args.$$$ "in") |-- Attrib.local_thm);
   457 
   458 fun inst_args f = f oo (#2 oo syntax insts);
   459 
   460 
   461 
   462 (** method text **)
   463 
   464 (* datatype text *)
   465 
   466 datatype text =
   467   Basic of (Proof.context -> Proof.method) |
   468   Source of Args.src |
   469   Then of text list |
   470   Orelse of text list |
   471   Try of text |
   472   Repeat1 of text;
   473 
   474 
   475 (* refine *)
   476 
   477 fun gen_refine f text state =
   478   let
   479     val thy = Proof.theory_of state;
   480 
   481     fun eval (Basic mth) = f mth
   482       | eval (Source src) = f (method thy src)
   483       | eval (Then txts) = Seq.EVERY (map eval txts)
   484       | eval (Orelse txts) = Seq.FIRST (map eval txts)
   485       | eval (Try txt) = Seq.TRY (eval txt)
   486       | eval (Repeat1 txt) = Seq.REPEAT1 (eval txt);
   487   in eval text state end;
   488 
   489 val refine = gen_refine Proof.refine;
   490 val refine_end = gen_refine Proof.refine_end;
   491 
   492 
   493 (* structured proof steps *)
   494 
   495 val default_text = Source (Args.src (("default", []), Position.none));
   496 val this_text = Basic (K this);
   497 
   498 fun finish ctxt = METHOD (K (FILTER Thm.no_prems (ALLGOALS (assm_tac ctxt) THEN flexflex_tac)));
   499 
   500 fun finish_text None = Basic finish
   501   | finish_text (Some txt) = Then [txt, Basic finish];
   502 
   503 fun proof opt_text state =
   504   state
   505   |> Proof.assert_backward
   506   |> refine (if_none opt_text default_text)
   507   |> Seq.map (Proof.goal_facts (K []))
   508   |> Seq.map Proof.enter_forward;
   509 
   510 fun local_qed opt_text = Proof.local_qed (refine (finish_text opt_text));
   511 fun local_terminal_proof (text, opt_text) pr = Seq.THEN (proof (Some text), local_qed opt_text pr);
   512 val local_immediate_proof = local_terminal_proof (this_text, None);
   513 val local_default_proof = local_terminal_proof (default_text, None);
   514 
   515 
   516 fun global_qeds opt_text = Proof.global_qed (refine (finish_text opt_text));
   517 
   518 fun global_qed opt_text state =
   519   state
   520   |> global_qeds opt_text
   521   |> Proof.check_result "Failed to finish proof" state
   522   |> Seq.hd;
   523 
   524 fun global_terminal_proof (text, opt_text) state =
   525   state
   526   |> proof (Some text)
   527   |> Proof.check_result "Terminal proof method failed" state
   528   |> (Seq.flat o Seq.map (global_qeds opt_text))
   529   |> Proof.check_result "Failed to finish proof (after successful terminal method)" state
   530   |> Seq.hd;
   531 
   532 val global_immediate_proof = global_terminal_proof (this_text, None);
   533 val global_default_proof = global_terminal_proof (default_text, None);
   534 
   535 
   536 
   537 (** theory setup **)
   538 
   539 (* pure_methods *)
   540 
   541 val pure_methods =
   542  [("fail", no_args fail, "force failure"),
   543   ("succeed", no_args succeed, "succeed"),
   544   ("-", no_args insert_facts, "do nothing, inserting current facts only"),
   545   ("insert", thms_args insert, "insert theorems, ignoring facts (improper!)"),
   546   ("unfold", thms_args unfold, "unfold definitions"),
   547   ("fold", thms_args fold, "fold definitions"),
   548   ("default", thms_ctxt_args some_rule, "apply some rule"),
   549   ("rule", thms_ctxt_args some_rule, "apply some rule"),
   550   ("erule", thms_ctxt_args some_erule, "apply some rule in elimination manner (improper!)"),
   551   ("drule", thms_ctxt_args some_drule, "apply some rule in destruct manner (improper!)"),
   552   ("frule", thms_ctxt_args some_frule, "apply some rule in forward manner (improper!)"),
   553   ("this", no_args this, "apply current facts as rules"),
   554   ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
   555   ("res_inst_tac", inst_args res_inst, "res_inst_tac emulation (beware!)"),
   556   ("eres_inst_tac", inst_args eres_inst, "eres_inst_tac emulation (beware!)"),
   557   ("dres_inst_tac", inst_args dres_inst, "dres_inst_tac emulation (beware!)"),
   558   ("forw_inst_tac", inst_args forw_inst, "forw_inst_tac emulation (beware!)"),
   559   ("prolog", thms_args prolog, "simple prolog interpreter")];
   560 
   561 
   562 (* setup *)
   563 
   564 val setup =
   565  [GlobalRules.init, LocalRules.init, Attrib.add_attributes rule_atts,
   566   MethodsData.init, add_methods pure_methods];
   567 
   568 
   569 end;
   570 
   571 
   572 structure BasicMethod: BASIC_METHOD = Method;
   573 open BasicMethod;