src/Pure/Isar/method.ML
author wenzelm
Mon, 02 May 2011 16:33:21 +0200
changeset 42616 92715b528e78
parent 42380 9371ea9f91fb
child 42813 6c841fa92fa2
permissions -rw-r--r--
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory; proper name bindings;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/method.ML
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
     3
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
     4
Isar proof methods.
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
     5
*)
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
     6
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
     7
signature BASIC_METHOD =
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
     8
sig
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
     9
  val FINDGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    10
  val HEADGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq
41379
b31d7a1cd08f configuration option "rule_trace";
wenzelm
parents: 39507
diff changeset
    11
  val rule_trace: bool Config.T
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    12
end;
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    13
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    14
signature METHOD =
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    15
sig
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    16
  include BASIC_METHOD
30508
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
    17
  type method
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
    18
  val apply: (Proof.context -> method) -> Proof.context -> thm list -> cases_tactic
18227
d4cfa0fee007 moved multi_resolve(s) to drule.ML;
wenzelm
parents: 18145
diff changeset
    19
  val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    20
  val RAW_METHOD: (thm list -> tactic) -> method
18227
d4cfa0fee007 moved multi_resolve(s) to drule.ML;
wenzelm
parents: 18145
diff changeset
    21
  val METHOD_CASES: (thm list -> cases_tactic) -> method
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    22
  val METHOD: (thm list -> tactic) -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    23
  val fail: method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    24
  val succeed: method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    25
  val insert_tac: thm list -> int -> tactic
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    26
  val insert: thm list -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    27
  val insert_facts: method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    28
  val SIMPLE_METHOD: tactic -> method
21592
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
    29
  val SIMPLE_METHOD': (int -> tactic) -> method
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
    30
  val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    31
  val defer: int option -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    32
  val prefer: int -> method
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    33
  val cheating: bool -> Proof.context -> method
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    34
  val intro: thm list -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    35
  val elim: thm list -> method
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    36
  val unfold: thm list -> Proof.context -> method
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    37
  val fold: thm list -> Proof.context -> method
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    38
  val atomize: bool -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    39
  val this: method
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    40
  val fact: thm list -> Proof.context -> method
30234
7dd251bce291 renamed Method.assumption_tac back to Method.assm_tac -- as assumption_tac it would have to be exactly the tactic behind the assumption method (with facts);
wenzelm
parents: 30190
diff changeset
    41
  val assm_tac: Proof.context -> int -> tactic
30567
cd8e20f86795 close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents: 30544
diff changeset
    42
  val all_assm_tac: Proof.context -> tactic
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    43
  val assumption: Proof.context -> method
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    44
  val close: bool -> Proof.context -> method
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    45
  val trace: Proof.context -> thm list -> unit
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5921
diff changeset
    46
  val rule_tac: thm list -> thm list -> int -> tactic
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    47
  val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic
25270
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
    48
  val intros_tac: thm list -> thm list -> tactic
36093
0880493627ca Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents: 33522
diff changeset
    49
  val try_intros_tac: thm list -> thm list -> tactic
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    50
  val rule: thm list -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    51
  val erule: int -> thm list -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    52
  val drule: int -> thm list -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    53
  val frule: int -> thm list -> method
27235
134991516430 ML tactic: do not abstract over context again;
wenzelm
parents: 26892
diff changeset
    54
  val set_tactic: (thm list -> tactic) -> Proof.context -> Proof.context
26385
ae7564661e76 ML runtime compilation: pass position, tuned signature;
wenzelm
parents: 26291
diff changeset
    55
  val tactic: string * Position.T -> Proof.context -> method
27235
134991516430 ML tactic: do not abstract over context again;
wenzelm
parents: 26892
diff changeset
    56
  val raw_tactic: string * Position.T -> Proof.context -> method
27729
aaf08262b177 tuned signature;
wenzelm
parents: 27383
diff changeset
    57
  type src = Args.src
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    58
  datatype text =
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
    59
    Basic of Proof.context -> method |
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15574
diff changeset
    60
    Source of src |
20030
e62913ef9d24 added method_i and Source_i;
wenzelm
parents: 19778
diff changeset
    61
    Source_i of src |
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    62
    Then of text list |
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    63
    Orelse of text list |
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    64
    Try of text |
19186
1bf4b5c4a794 text: added SelectGoals;
wenzelm
parents: 19046
diff changeset
    65
    Repeat1 of text |
1bf4b5c4a794 text: added SelectGoals;
wenzelm
parents: 19046
diff changeset
    66
    SelectGoals of int * text
17857
810a67ecbc64 added primitive_text, succeed_text;
wenzelm
parents: 17756
diff changeset
    67
  val primitive_text: (thm -> thm) -> text
810a67ecbc64 added primitive_text, succeed_text;
wenzelm
parents: 17756
diff changeset
    68
  val succeed_text: text
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    69
  val default_text: text
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    70
  val this_text: text
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    71
  val done_text: text
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
    72
  val sorry_text: bool -> text
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
    73
  val finish_text: text option * bool -> text
27729
aaf08262b177 tuned signature;
wenzelm
parents: 27383
diff changeset
    74
  val print_methods: theory -> unit
26892
9454a8bd1114 added intern, defined;
wenzelm
parents: 26762
diff changeset
    75
  val intern: theory -> xstring -> string
9454a8bd1114 added intern, defined;
wenzelm
parents: 26762
diff changeset
    76
  val defined: theory -> string -> bool
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    77
  val method: theory -> src -> Proof.context -> method
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    78
  val method_i: theory -> src -> Proof.context -> method
30512
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
    79
  val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
    80
  val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory
30575
368e26dfba69 more precise type Symbol_Pos.text;
wenzelm
parents: 30567
diff changeset
    81
  val method_setup: bstring * Position.T -> Symbol_Pos.text * Position.T -> string ->
368e26dfba69 more precise type Symbol_Pos.text;
wenzelm
parents: 30567
diff changeset
    82
    theory -> theory
30540
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
    83
  type modifier = (Proof.context -> Proof.context) * attribute
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
    84
  val section: modifier parser list -> thm list context_parser
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
    85
  val sections: modifier parser list -> thm list list context_parser
30512
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
    86
  val parse: text parser
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    87
end;
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    88
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    89
structure Method: METHOD =
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    90
struct
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    91
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    92
(** generic tools **)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    93
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    94
(* goal addressing *)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    95
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    96
fun FINDGOAL tac st =
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    97
  let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    98
  in find 1 (Thm.nprems_of st) st end;
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    99
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   100
fun HEADGOAL tac = tac 1;
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   101
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   102
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   103
12324
5db4b4596d1a rule context and attributes moved to rule_context.ML;
wenzelm
parents: 12311
diff changeset
   104
(** proof methods **)
5db4b4596d1a rule context and attributes moved to rule_context.ML;
wenzelm
parents: 12311
diff changeset
   105
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   106
(* datatype method *)
11731
1a0c1ef86518 added trace_rules, trace;
wenzelm
parents: 10907
diff changeset
   107
18227
d4cfa0fee007 moved multi_resolve(s) to drule.ML;
wenzelm
parents: 18145
diff changeset
   108
datatype method = Meth of thm list -> cases_tactic;
11731
1a0c1ef86518 added trace_rules, trace;
wenzelm
parents: 10907
diff changeset
   109
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   110
fun apply meth ctxt = let val Meth m = meth ctxt in m end;
11731
1a0c1ef86518 added trace_rules, trace;
wenzelm
parents: 10907
diff changeset
   111
17756
d4a35f82fbb4 minor tweaks for Poplog/ML;
wenzelm
parents: 17587
diff changeset
   112
val RAW_METHOD_CASES = Meth;
11731
1a0c1ef86518 added trace_rules, trace;
wenzelm
parents: 10907
diff changeset
   113
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   114
fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac);
12144
f84eb7334d04 added RAW_METHOD, RAW_METHOD_CASES;
wenzelm
parents: 12119
diff changeset
   115
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   116
fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts =>
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21592
diff changeset
   117
  Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts));
8372
7b2cec1e789c added METHOD_CASES, resolveq_cases_tac;
wenzelm
parents: 8351
diff changeset
   118
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21592
diff changeset
   119
fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   120
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   121
val fail = METHOD (K no_tac);
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   122
val succeed = METHOD (K all_tac);
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   123
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   124
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   125
(* insert facts *)
7419
wenzelm
parents: 7367
diff changeset
   126
wenzelm
parents: 7367
diff changeset
   127
local
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   128
21579
abd2b4386a63 COMP_INCR;
wenzelm
parents: 20335
diff changeset
   129
fun cut_rule_tac rule =
abd2b4386a63 COMP_INCR;
wenzelm
parents: 20335
diff changeset
   130
  Tactic.rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl);
6981
eaade7e398a7 export assumption_tac;
wenzelm
parents: 6951
diff changeset
   131
7419
wenzelm
parents: 7367
diff changeset
   132
in
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   133
7419
wenzelm
parents: 7367
diff changeset
   134
fun insert_tac [] i = all_tac
wenzelm
parents: 7367
diff changeset
   135
  | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts);
6981
eaade7e398a7 export assumption_tac;
wenzelm
parents: 6951
diff changeset
   136
7555
dd281afb33d7 setup for refined facts handling;
wenzelm
parents: 7526
diff changeset
   137
val insert_facts = METHOD (ALLGOALS o insert_tac);
7664
c151ac595551 insert: ignore facts;
wenzelm
parents: 7611
diff changeset
   138
fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms));
7419
wenzelm
parents: 7367
diff changeset
   139
9706
8e48a19fc81e removed METHOD0;
wenzelm
parents: 9653
diff changeset
   140
fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac);
21592
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
   141
fun SIMPLE_METHOD'' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac));
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
   142
val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL;
9706
8e48a19fc81e removed METHOD0;
wenzelm
parents: 9653
diff changeset
   143
12324
5db4b4596d1a rule context and attributes moved to rule_context.ML;
wenzelm
parents: 12311
diff changeset
   144
end;
5db4b4596d1a rule context and attributes moved to rule_context.ML;
wenzelm
parents: 12311
diff changeset
   145
9706
8e48a19fc81e removed METHOD0;
wenzelm
parents: 9653
diff changeset
   146
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   147
(* shuffle subgoals *)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   148
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   149
fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)));
18939
wenzelm
parents: 18921
diff changeset
   150
fun defer opt_i = METHOD (K (Tactic.defer_tac (the_default 1 opt_i)));
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   151
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   152
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   153
(* cheating *)
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   154
32969
15489e162b21 Method.cheating: check quick_and_dirty here;
wenzelm
parents: 32966
diff changeset
   155
fun cheating int ctxt =
15489e162b21 Method.cheating: check quick_and_dirty here;
wenzelm
parents: 32966
diff changeset
   156
  if int orelse ! quick_and_dirty then
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   157
    METHOD (K (Skip_Proof.cheat_tac (Proof_Context.theory_of ctxt)))
32969
15489e162b21 Method.cheating: check quick_and_dirty here;
wenzelm
parents: 32966
diff changeset
   158
  else error "Cheating requires quick_and_dirty mode!";
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   159
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   160
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   161
(* unfold intro/elim rules *)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   162
21592
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
   163
fun intro ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths));
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
   164
fun elim ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths));
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   165
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   166
12384
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   167
(* unfold/fold definitions *)
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   168
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 33522
diff changeset
   169
fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths));
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 33522
diff changeset
   170
fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths));
6532
9d79a304aecc fold / unfold methods;
wenzelm
parents: 6500
diff changeset
   171
12384
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   172
12829
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   173
(* atomize rule statements *)
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   174
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35624
diff changeset
   175
fun atomize false = SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac)
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35624
diff changeset
   176
  | atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac)));
12829
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   177
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   178
18039
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   179
(* this -- resolve facts directly *)
12384
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   180
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   181
val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac));
9484
3bda55143260 (un)fold: CHANGED;
wenzelm
parents: 9418
diff changeset
   182
3bda55143260 (un)fold: CHANGED;
wenzelm
parents: 9418
diff changeset
   183
18039
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   184
(* fact -- composition by facts from context *)
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   185
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   186
fun fact [] ctxt = SIMPLE_METHOD' (Proof_Context.some_fact_tac ctxt)
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   187
  | fact rules _ = SIMPLE_METHOD' (Proof_Context.fact_tac rules);
18039
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   188
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   189
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   190
(* assumption *)
7419
wenzelm
parents: 7367
diff changeset
   191
wenzelm
parents: 7367
diff changeset
   192
local
wenzelm
parents: 7367
diff changeset
   193
19778
f0a318495ca4 assm_tac: try rule termI;
wenzelm
parents: 19482
diff changeset
   194
fun cond_rtac cond rule = SUBGOAL (fn (prop, i) =>
f0a318495ca4 assm_tac: try rule termI;
wenzelm
parents: 19482
diff changeset
   195
  if cond (Logic.strip_assums_concl prop)
f0a318495ca4 assm_tac: try rule termI;
wenzelm
parents: 19482
diff changeset
   196
  then Tactic.rtac rule i else no_tac);
7419
wenzelm
parents: 7367
diff changeset
   197
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29301
diff changeset
   198
in
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29301
diff changeset
   199
30234
7dd251bce291 renamed Method.assumption_tac back to Method.assm_tac -- as assumption_tac it would have to be exactly the tactic behind the assumption method (with facts);
wenzelm
parents: 30190
diff changeset
   200
fun assm_tac ctxt =
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   201
  assume_tac APPEND'
23349
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   202
  Goal.assume_rule_tac ctxt APPEND'
19778
f0a318495ca4 assm_tac: try rule termI;
wenzelm
parents: 19482
diff changeset
   203
  cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND'
f0a318495ca4 assm_tac: try rule termI;
wenzelm
parents: 19482
diff changeset
   204
  cond_rtac (can Logic.dest_term) Drule.termI;
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   205
30567
cd8e20f86795 close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents: 30544
diff changeset
   206
fun all_assm_tac ctxt st = EVERY1 (replicate (Thm.nprems_of st) (assm_tac ctxt)) st;
cd8e20f86795 close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents: 30544
diff changeset
   207
23349
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   208
fun assumption ctxt = METHOD (HEADGOAL o
30234
7dd251bce291 renamed Method.assumption_tac back to Method.assm_tac -- as assumption_tac it would have to be exactly the tactic behind the assumption method (with facts);
wenzelm
parents: 30190
diff changeset
   209
  (fn [] => assm_tac ctxt
23349
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   210
    | [fact] => solve_tac [fact]
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   211
    | _ => K no_tac));
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   212
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   213
fun close immed ctxt = METHOD (K
30567
cd8e20f86795 close/all_assm_tac: finish all subgoals from left to right (cf. Proof.goal_tac a28d83e903ce) -- NB: ALLGOALS/THEN_ALL_NEW operate from right to left;
wenzelm
parents: 30544
diff changeset
   214
  (FILTER Thm.no_prems ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac)));
7419
wenzelm
parents: 7367
diff changeset
   215
wenzelm
parents: 7367
diff changeset
   216
end;
wenzelm
parents: 7367
diff changeset
   217
wenzelm
parents: 7367
diff changeset
   218
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   219
(* rule etc. -- single-step refinements *)
12347
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   220
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42380
diff changeset
   221
val rule_trace = Attrib.setup_config_bool (Binding.name "rule_trace") (fn _ => false);
12347
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   222
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   223
fun trace ctxt rules =
41379
b31d7a1cd08f configuration option "rule_trace";
wenzelm
parents: 39507
diff changeset
   224
  if Config.get ctxt rule_trace andalso not (null rules) then
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31304
diff changeset
   225
    Pretty.big_list "rules:" (map (Display.pretty_thm ctxt) rules)
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21879
diff changeset
   226
    |> Pretty.string_of |> tracing
279b129498b6 removed conditional combinator;
wenzelm
parents: 21879
diff changeset
   227
  else ();
12347
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   228
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   229
local
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   230
18841
edecd40194c1 method (un)folded: option '(raw)';
wenzelm
parents: 18824
diff changeset
   231
fun gen_rule_tac tac rules facts =
edecd40194c1 method (un)folded: option '(raw)';
wenzelm
parents: 18824
diff changeset
   232
  (fn i => fn st =>
edecd40194c1 method (un)folded: option '(raw)';
wenzelm
parents: 18824
diff changeset
   233
    if null facts then tac rules i st
edecd40194c1 method (un)folded: option '(raw)';
wenzelm
parents: 18824
diff changeset
   234
    else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules))
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21592
diff changeset
   235
  THEN_ALL_NEW Goal.norm_hhf_tac;
7130
a17f7b5ac40f added erule;
wenzelm
parents: 6981
diff changeset
   236
10744
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   237
fun gen_arule_tac tac j rules facts =
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   238
  EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac);
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   239
11785
3087d6f19adc intro! and elim! rules;
wenzelm
parents: 11765
diff changeset
   240
fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) =>
3087d6f19adc intro! and elim! rules;
wenzelm
parents: 11765
diff changeset
   241
  let
3087d6f19adc intro! and elim! rules;
wenzelm
parents: 11765
diff changeset
   242
    val rules =
3087d6f19adc intro! and elim! rules;
wenzelm
parents: 11765
diff changeset
   243
      if not (null arg_rules) then arg_rules
33369
470a7b233ee5 modernized structure Context_Rules;
wenzelm
parents: 33159
diff changeset
   244
      else flat (Context_Rules.find_rules false facts goal ctxt)
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12007
diff changeset
   245
  in trace ctxt rules; tac rules facts i end);
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10034
diff changeset
   246
10744
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   247
fun meth tac x = METHOD (HEADGOAL o tac x);
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   248
fun meth' tac x y = METHOD (HEADGOAL o tac x y);
8220
e04928747b18 [df]rule methods;
wenzelm
parents: 8205
diff changeset
   249
7419
wenzelm
parents: 7367
diff changeset
   250
in
wenzelm
parents: 7367
diff changeset
   251
10744
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   252
val rule_tac = gen_rule_tac Tactic.resolve_tac;
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   253
val rule = meth rule_tac;
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   254
val some_rule_tac = gen_some_rule_tac rule_tac;
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   255
val some_rule = meth' some_rule_tac;
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   256
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   257
val erule = meth' (gen_arule_tac Tactic.eresolve_tac);
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   258
val drule = meth' (gen_arule_tac Tactic.dresolve_tac);
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   259
val frule = meth' (gen_arule_tac Tactic.forward_tac);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   260
7419
wenzelm
parents: 7367
diff changeset
   261
end;
wenzelm
parents: 7367
diff changeset
   262
wenzelm
parents: 7367
diff changeset
   263
25270
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   264
(* intros_tac -- pervasive search spanned by intro rules *)
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   265
36093
0880493627ca Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents: 33522
diff changeset
   266
fun gen_intros_tac goals intros facts =
0880493627ca Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents: 33522
diff changeset
   267
  goals (insert_tac facts THEN'
25270
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   268
      REPEAT_ALL_NEW (resolve_tac intros))
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   269
    THEN Tactic.distinct_subgoals_tac;
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   270
36093
0880493627ca Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents: 33522
diff changeset
   271
val intros_tac = gen_intros_tac ALLGOALS;
0880493627ca Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents: 33522
diff changeset
   272
val try_intros_tac = gen_intros_tac TRYALL;
25270
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   273
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
   274
8351
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   275
(* ML tactics *)
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   276
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
   277
structure ML_Tactic = Proof_Data
26472
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   278
(
27235
134991516430 ML tactic: do not abstract over context again;
wenzelm
parents: 26892
diff changeset
   279
  type T = thm list -> tactic;
26472
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   280
  fun init _ = undefined;
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   281
);
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   282
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
   283
val set_tactic = ML_Tactic.put;
8351
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   284
26472
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   285
fun ml_tactic (txt, pos) ctxt =
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   286
  let
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   287
    val ctxt' = ctxt |> Context.proof_map
27235
134991516430 ML tactic: do not abstract over context again;
wenzelm
parents: 26892
diff changeset
   288
      (ML_Context.expression pos
134991516430 ML tactic: do not abstract over context again;
wenzelm
parents: 26892
diff changeset
   289
        "fun tactic (facts: thm list) : tactic"
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   290
        "Context.map_proof (Method.set_tactic tactic)" (ML_Lex.read pos txt));
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
   291
  in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (ML_Tactic.get ctxt') end;
23425
b74315510f85 added raw_tactic;
wenzelm
parents: 23395
diff changeset
   292
b74315510f85 added raw_tactic;
wenzelm
parents: 23395
diff changeset
   293
fun tactic txt ctxt = METHOD (ml_tactic txt ctxt);
b74315510f85 added raw_tactic;
wenzelm
parents: 23395
diff changeset
   294
fun raw_tactic txt ctxt = RAW_METHOD (ml_tactic txt ctxt);
8351
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   295
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   296
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   297
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   298
(** method syntax **)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   299
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   300
(* method text *)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   301
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   302
type src = Args.src;
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   303
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   304
datatype text =
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   305
  Basic of Proof.context -> method |
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   306
  Source of src |
20030
e62913ef9d24 added method_i and Source_i;
wenzelm
parents: 19778
diff changeset
   307
  Source_i of src |
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   308
  Then of text list |
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   309
  Orelse of text list |
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   310
  Try of text |
19186
1bf4b5c4a794 text: added SelectGoals;
wenzelm
parents: 19046
diff changeset
   311
  Repeat1 of text |
1bf4b5c4a794 text: added SelectGoals;
wenzelm
parents: 19046
diff changeset
   312
  SelectGoals of int * text;
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   313
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   314
fun primitive_text r = Basic (K (SIMPLE_METHOD (PRIMITIVE r)));
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   315
val succeed_text = Basic (K succeed);
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   316
val default_text = Source (Args.src (("default", []), Position.none));
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   317
val this_text = Basic (K this);
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   318
val done_text = Basic (K (SIMPLE_METHOD all_tac));
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   319
fun sorry_text int = Basic (cheating int);
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   320
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   321
fun finish_text (NONE, immed) = Basic (close immed)
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   322
  | finish_text (SOME txt, immed) = Then [txt, Basic (close immed)];
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   323
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   324
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   325
(* method definitions *)
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   326
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   327
structure Methods = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   328
(
33095
bbd52d2f8696 renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents: 33092
diff changeset
   329
  type T = ((src -> Proof.context -> method) * string) Name_Space.table;
33159
369da293bbd4 make SML/NJ happy;
wenzelm
parents: 33096
diff changeset
   330
  val empty : T = Name_Space.empty_table "method";
16448
6c45c5416b79 (RAW_)METHOD_CASES: RuleCases.tactic;
wenzelm
parents: 16347
diff changeset
   331
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33519
diff changeset
   332
  fun merge data : T = Name_Space.merge_tables data;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   333
);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   334
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   335
fun print_methods thy =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   336
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   337
    val ctxt = Proof_Context.init_global thy;
24116
wenzelm
parents: 24022
diff changeset
   338
    val meths = Methods.get thy;
33092
c859019d3ac5 eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
wenzelm
parents: 32969
diff changeset
   339
    fun prt_meth (name, (_, comment)) = Pretty.block
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   340
      [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   341
  in
42358
b47d41d9f4b5 Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents: 41379
diff changeset
   342
    [Pretty.big_list "methods:" (map prt_meth (Name_Space.extern_table ctxt meths))]
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   343
    |> Pretty.chunks |> Pretty.writeln
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   344
  end;
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7601
diff changeset
   345
33092
c859019d3ac5 eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
wenzelm
parents: 32969
diff changeset
   346
fun add_method name meth comment thy = thy
42375
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42360
diff changeset
   347
  |> Methods.map
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42360
diff changeset
   348
    (Name_Space.define (Proof_Context.init_global thy) true (Sign.naming_of thy)
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42360
diff changeset
   349
      (name, (meth, comment)) #> snd);
31304
00a9c674cf40 eliminated old Method.add_method(s);
wenzelm
parents: 31303
diff changeset
   350
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   351
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   352
(* get methods *)
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   353
33095
bbd52d2f8696 renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents: 33092
diff changeset
   354
val intern = Name_Space.intern o #1 o Methods.get;
26892
9454a8bd1114 added intern, defined;
wenzelm
parents: 26762
diff changeset
   355
val defined = Symtab.defined o #2 o Methods.get;
9454a8bd1114 added intern, defined;
wenzelm
parents: 26762
diff changeset
   356
20030
e62913ef9d24 added method_i and Source_i;
wenzelm
parents: 19778
diff changeset
   357
fun method_i thy =
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   358
  let
42380
9371ea9f91fb markup attributes/methods via name space;
wenzelm
parents: 42375
diff changeset
   359
    val (space, tab) = Methods.get thy;
5884
113badd4dae5 several args parsers;
wenzelm
parents: 5824
diff changeset
   360
    fun meth src =
20030
e62913ef9d24 added method_i and Source_i;
wenzelm
parents: 19778
diff changeset
   361
      let val ((name, _), pos) = Args.dest_src src in
42380
9371ea9f91fb markup attributes/methods via name space;
wenzelm
parents: 42375
diff changeset
   362
        (case Symtab.lookup tab name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   363
          NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
42380
9371ea9f91fb markup attributes/methods via name space;
wenzelm
parents: 42375
diff changeset
   364
        | SOME (mth, _) => (Position.report pos (Name_Space.markup space name); mth src))
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   365
      end;
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   366
  in meth end;
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   367
33095
bbd52d2f8696 renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents: 33092
diff changeset
   368
fun method thy = method_i thy o Args.map_name (Name_Space.intern (#1 (Methods.get thy)));
20030
e62913ef9d24 added method_i and Source_i;
wenzelm
parents: 19778
diff changeset
   369
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   370
30512
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
   371
(* method setup *)
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
   372
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
   373
fun syntax scan = Args.context_syntax "method" scan;
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
   374
31304
00a9c674cf40 eliminated old Method.add_method(s);
wenzelm
parents: 31303
diff changeset
   375
fun setup name scan =
00a9c674cf40 eliminated old Method.add_method(s);
wenzelm
parents: 31303
diff changeset
   376
  add_method name
00a9c674cf40 eliminated old Method.add_method(s);
wenzelm
parents: 31303
diff changeset
   377
    (fn src => fn ctxt => let val (m, ctxt') = syntax scan src ctxt in m ctxt' end);
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   378
26385
ae7564661e76 ML runtime compilation: pass position, tuned signature;
wenzelm
parents: 26291
diff changeset
   379
fun method_setup name (txt, pos) cmt =
26472
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   380
  Context.theory_map (ML_Context.expression pos
30544
0ed8fe16331a adapted 'method_setup' command to Method.setup;
wenzelm
parents: 30540
diff changeset
   381
    "val (name, scan, comment): binding * (Proof.context -> Proof.method) context_parser * string"
0ed8fe16331a adapted 'method_setup' command to Method.setup;
wenzelm
parents: 30540
diff changeset
   382
    "Context.map_theory (Method.setup name scan comment)"
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   383
    (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   384
      ML_Lex.read pos txt @
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   385
      ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")")));
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   386
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   387
5884
113badd4dae5 several args parsers;
wenzelm
parents: 5824
diff changeset
   388
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   389
(** concrete syntax **)
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   390
5884
113badd4dae5 several args parsers;
wenzelm
parents: 5824
diff changeset
   391
(* sections *)
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   392
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
   393
type modifier = (Proof.context -> Proof.context) * attribute;
7268
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   394
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   395
local
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   396
24010
2ef318813e1a method section scanners: added [[declaration]] syntax, ignore sid-effects of thms;
wenzelm
parents: 23937
diff changeset
   397
fun thms ss = Scan.repeat (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm) >> flat;
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 30165
diff changeset
   398
fun app (f, att) (context, ths) = Library.foldl_map att (Context.map_proof f context, ths);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   399
30540
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
   400
in
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
   401
24022
wenzelm
parents: 24010
diff changeset
   402
fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|--
wenzelm
parents: 24010
diff changeset
   403
  (fn (m, ths) => Scan.succeed (app m (context, ths))));
5884
113badd4dae5 several args parsers;
wenzelm
parents: 5824
diff changeset
   404
30540
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
   405
fun sections ss = Scan.repeat (section ss);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   406
7268
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   407
end;
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   408
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   409
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   410
(* extra rule methods *)
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   411
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   412
fun xrule_meth m =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   413
  Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >>
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   414
  (fn (n, ths) => K (m n ths));
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   415
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   416
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   417
(* outer parser *)
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   418
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   419
fun is_symid_meth s =
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   420
  s <> "|" andalso s <> "?" andalso s <> "+" andalso Token.ident_or_symbolic s;
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   421
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   422
local
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   423
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   424
fun meth4 x =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   425
 (Parse.position (Parse.xname >> rpair []) >> (Source o Args.src) ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   426
  Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   427
and meth3 x =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   428
 (meth4 --| Parse.$$$ "?" >> Try ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   429
  meth4 --| Parse.$$$ "+" >> Repeat1 ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   430
  meth4 -- (Parse.$$$ "[" |-- Scan.optional Parse.nat 1 --| Parse.$$$ "]") >> (SelectGoals o swap) ||
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   431
  meth4) x
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   432
and meth2 x =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   433
 (Parse.position (Parse.xname -- Args.parse1 is_symid_meth) >> (Source o Args.src) ||
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   434
  meth3) x
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   435
and meth1 x = (Parse.enum1 "," meth2 >> (fn [m] => m | ms => Then ms)) x
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   436
and meth0 x = (Parse.enum1 "|" meth1 >> (fn [m] => m | ms => Orelse ms)) x;
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   437
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   438
in val parse = meth3 end;
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   439
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   440
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18640
diff changeset
   441
(* theory setup *)
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   442
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26455
diff changeset
   443
val _ = Context.>> (Context.map_theory
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42380
diff changeset
   444
 (setup (Binding.name "fail") (Scan.succeed (K fail)) "force failure" #>
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   445
  setup (Binding.name "succeed") (Scan.succeed (K succeed)) "succeed" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   446
  setup (Binding.name "-") (Scan.succeed (K insert_facts))
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   447
    "do nothing (insert current facts only)" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   448
  setup (Binding.name "insert") (Attrib.thms >> (K o insert))
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   449
    "insert theorems, ignoring facts (improper)" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   450
  setup (Binding.name "intro") (Attrib.thms >> (K o intro))
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   451
    "repeatedly apply introduction rules" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   452
  setup (Binding.name "elim") (Attrib.thms >> (K o elim))
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   453
    "repeatedly apply elimination rules" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   454
  setup (Binding.name "unfold") (Attrib.thms >> unfold_meth) "unfold definitions" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   455
  setup (Binding.name "fold") (Attrib.thms >> fold_meth) "fold definitions" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   456
  setup (Binding.name "atomize") (Args.mode "full" >> (K o atomize))
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   457
    "present local premises as object-level statements" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   458
  setup (Binding.name "rule") (Attrib.thms >> some_rule) "apply some intro/elim rule" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   459
  setup (Binding.name "erule") (xrule_meth erule) "apply rule in elimination manner (improper)" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   460
  setup (Binding.name "drule") (xrule_meth drule) "apply rule in destruct manner (improper)" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   461
  setup (Binding.name "frule") (xrule_meth frule) "apply rule in forward manner (improper)" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   462
  setup (Binding.name "this") (Scan.succeed (K this)) "apply current facts as rules" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   463
  setup (Binding.name "fact") (Attrib.thms >> fact) "composition by facts from context" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   464
  setup (Binding.name "assumption") (Scan.succeed assumption)
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   465
    "proof by assumption, preferring facts" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   466
  setup (Binding.name "rename_tac") (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >>
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   467
    (fn (quant, xs) => K (SIMPLE_METHOD'' quant (Tactic.rename_tac xs))))
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   468
    "rename parameters of goal" #>
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   469
  setup (Binding.name "rotate_tac") (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >>
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   470
    (fn (quant, i) => K (SIMPLE_METHOD'' quant (Tactic.rotate_tac i))))
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   471
      "rotate assumptions of goal" #>
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   472
  setup (Binding.name "tactic") (Scan.lift Args.name_source_position >> tactic)
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   473
    "ML tactic as proof method" #>
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 36959
diff changeset
   474
  setup (Binding.name "raw_tactic") (Scan.lift Args.name_source_position >> raw_tactic)
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   475
    "ML tactic as raw proof method"));
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   476
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   477
16145
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   478
(*final declarations of this structure!*)
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   479
val unfold = unfold_meth;
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   480
val fold = fold_meth;
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   481
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   482
end;
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   483
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   484
structure Basic_Method: BASIC_METHOD = Method;
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   485
open Basic_Method;
30508
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   486
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   487
val RAW_METHOD_CASES = Method.RAW_METHOD_CASES;
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   488
val RAW_METHOD = Method.RAW_METHOD;
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   489
val METHOD_CASES = Method.METHOD_CASES;
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   490
val METHOD = Method.METHOD;
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   491
val SIMPLE_METHOD = Method.SIMPLE_METHOD;
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   492
val SIMPLE_METHOD' = Method.SIMPLE_METHOD';
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   493
val SIMPLE_METHOD'' = Method.SIMPLE_METHOD'';
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   494