src/Pure/Isar/method.ML
author wenzelm
Thu, 25 Oct 2018 21:29:08 +0200
changeset 69187 d8849cfad60f
parent 68823 5e7b1ae10eb8
child 69216 1a52baa70aed
permissions -rw-r--r--
clarified ML position: proper markup for def/ref scopes (see also 162a4c2e97bc);
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
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
     7
infix 1 CONTEXT_THEN_ALL_NEW;
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
     8
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
     9
signature BASIC_METHOD =
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
    10
sig
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    11
  type context_state = Proof.context * thm
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    12
  type context_tactic = context_state -> context_state Seq.result Seq.seq
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    13
  val CONTEXT_CASES: Rule_Cases.cases -> tactic -> context_tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    14
  val CONTEXT_SUBGOAL: (term * int -> context_tactic) -> int -> context_tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    15
  val CONTEXT_THEN_ALL_NEW: (int -> context_tactic) * (int -> tactic) -> int -> context_tactic
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
    16
end;
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
    17
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    18
signature METHOD =
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    19
sig
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
    20
  include BASIC_METHOD
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    21
  type method = thm list -> context_tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    22
  val CONTEXT: Proof.context -> thm Seq.seq -> context_state Seq.result Seq.seq
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    23
  val CONTEXT_TACTIC: tactic -> context_tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    24
  val NO_CONTEXT_TACTIC: Proof.context -> context_tactic -> tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    25
  val CONTEXT_METHOD: (thm list -> context_tactic) -> method
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    26
  val METHOD: (thm list -> tactic) -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    27
  val fail: method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    28
  val succeed: method
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    29
  val insert_tac: Proof.context -> thm list -> int -> tactic
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    30
  val insert: thm list -> method
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    31
  val SIMPLE_METHOD: tactic -> method
21592
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
    32
  val SIMPLE_METHOD': (int -> tactic) -> method
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
    33
  val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    34
  val goal_cases_tac: string list -> context_tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    35
  val cheating: bool -> method
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
    36
  val intro: Proof.context -> thm list -> method
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
    37
  val elim: Proof.context -> thm list -> method
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    38
  val unfold: thm list -> Proof.context -> method
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    39
  val fold: thm list -> Proof.context -> method
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
    40
  val atomize: bool -> Proof.context -> method
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
    41
  val this: Proof.context -> method
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    42
  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
    43
  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
    44
  val all_assm_tac: Proof.context -> tactic
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    45
  val assumption: Proof.context -> method
46466
61c7214b4885 tuned signature, according to actual usage of these operations;
wenzelm
parents: 45375
diff changeset
    46
  val rule_trace: bool Config.T
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    47
  val trace: Proof.context -> thm list -> unit
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
    48
  val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
    49
  val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
    50
  val intros_tac: Proof.context -> thm list -> thm list -> tactic
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
    51
  val try_intros_tac: Proof.context -> thm list -> thm list -> tactic
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
    52
  val rule: Proof.context -> thm list -> method
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
    53
  val erule: Proof.context -> int -> thm list -> method
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
    54
  val drule: Proof.context -> int -> thm list -> method
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
    55
  val frule: Proof.context -> int -> thm list -> method
58018
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
    56
  val set_tactic: (morphism -> thm list -> tactic) -> Context.generic -> Context.generic
63251
9a20078425b1 clean facts more uniformly;
wenzelm
parents: 63250
diff changeset
    57
  val clean_facts: thm list -> thm list
63256
wenzelm
parents: 63252
diff changeset
    58
  val set_facts: thm list -> Proof.context -> Proof.context
wenzelm
parents: 63252
diff changeset
    59
  val get_facts: Proof.context -> thm list
55765
ec7ca5388dea markup for method combinators;
wenzelm
parents: 55761
diff changeset
    60
  type combinator_info
ec7ca5388dea markup for method combinators;
wenzelm
parents: 55761
diff changeset
    61
  val no_combinator_info: combinator_info
59660
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
    62
  datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
    63
  datatype text =
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
    64
    Source of Token.src |
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
    65
    Basic of Proof.context -> method |
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
    66
    Combinator of combinator_info * combinator * text list
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
    67
  val map_source: (Token.src -> Token.src) -> text -> text
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54742
diff changeset
    68
  val primitive_text: (Proof.context -> thm -> thm) -> text
17857
810a67ecbc64 added primitive_text, succeed_text;
wenzelm
parents: 17756
diff changeset
    69
  val succeed_text: text
60618
4c79543cc376 renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents: 60609
diff changeset
    70
  val standard_text: text
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    71
  val this_text: text
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
    72
  val done_text: text
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
    73
  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
    74
  val finish_text: text option * bool -> text
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59914
diff changeset
    75
  val print_methods: bool -> Proof.context -> unit
55742
a989bdaf8121 modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents: 55709
diff changeset
    76
  val check_name: Proof.context -> xstring * Position.T -> string
59907
6c0f62490699 tuned signature;
wenzelm
parents: 59666
diff changeset
    77
  val check_src: Proof.context -> Token.src -> Token.src
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
    78
  val check_text: Proof.context -> text -> text
61917
35ec3757d3c1 check and report source at most once, notably in body of "match" method;
wenzelm
parents: 61843
diff changeset
    79
  val checked_text: text -> bool
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
    80
  val method_syntax: (Proof.context -> method) context_parser ->
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
    81
    Token.src -> Proof.context -> method
30512
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
    82
  val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory
57935
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
    83
  val local_setup: binding -> (Proof.context -> method) context_parser -> string ->
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
    84
    local_theory -> local_theory
59064
a8bcb5a446c8 more abstract type Input.source;
wenzelm
parents: 58991
diff changeset
    85
  val method_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
    86
  val method: Proof.context -> Token.src -> Proof.context -> method
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
    87
  val method_closure: Proof.context -> Token.src -> Token.src
59909
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
    88
  val closure: bool Config.T
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
    89
  val method_cmd: Proof.context -> Token.src -> Proof.context -> method
60553
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
    90
  val detect_closure_state: thm -> bool
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    91
  val STATIC: (unit -> unit) -> context_tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    92
  val RUNTIME: context_tactic -> context_tactic
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
    93
  val sleep: Time.time -> context_tactic
58007
671c607fb4af just one context for Method.evaluate (in contrast to a989bdaf8121, but in accordance to old global situation);
wenzelm
parents: 58006
diff changeset
    94
  val evaluate: text -> Proof.context -> method
63527
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
    95
  val evaluate_runtime: text -> Proof.context -> method
58048
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
    96
  type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T}
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
    97
  val modifier: attribute -> Position.T -> modifier
59982
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
    98
  val old_section_parser: bool Config.T
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
    99
  val sections: modifier parser list -> unit context_parser
49889
00ea087e83d8 more method position information, notably finished_pos after end of previous text;
wenzelm
parents: 49866
diff changeset
   100
  type text_range = text * Position.range
00ea087e83d8 more method position information, notably finished_pos after end of previous text;
wenzelm
parents: 49866
diff changeset
   101
  val text: text_range option -> text option
00ea087e83d8 more method position information, notably finished_pos after end of previous text;
wenzelm
parents: 49866
diff changeset
   102
  val position: text_range option -> Position.T
55795
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   103
  val reports_of: text_range -> Position.report list
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   104
  val report: text_range -> unit
59981
0d0f9c66ff3f tuned signature;
wenzelm
parents: 59953
diff changeset
   105
  val parser: int -> text_range parser
55761
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   106
  val parse: text_range parser
63527
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   107
  val read: Proof.context -> Token.src -> text
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   108
  val read_closure: Proof.context -> Token.src -> text * Token.src
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   109
  val read_closure_input: Proof.context -> Input.source -> text * Token.src
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   110
  val text_closure: text context_parser
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   111
end;
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   112
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   113
structure Method: METHOD =
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   114
struct
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   115
12324
5db4b4596d1a rule context and attributes moved to rule_context.ML;
wenzelm
parents: 12311
diff changeset
   116
(** proof methods **)
5db4b4596d1a rule context and attributes moved to rule_context.ML;
wenzelm
parents: 12311
diff changeset
   117
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   118
(* tactics with proof context / cases *)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   119
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   120
type context_state = Proof.context * thm;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   121
type context_tactic = context_state -> context_state Seq.result Seq.seq;
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   122
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   123
fun CONTEXT ctxt : thm Seq.seq -> context_state Seq.result Seq.seq =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   124
  Seq.map (Seq.Result o pair ctxt);
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   125
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   126
fun CONTEXT_TACTIC tac : context_tactic =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   127
  fn (ctxt, st) => CONTEXT ctxt (tac st);
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   128
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   129
fun NO_CONTEXT_TACTIC ctxt (tac: context_tactic) st =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   130
  tac (ctxt, st) |> Seq.filter_results |> Seq.map snd;
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   131
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   132
fun CONTEXT_CASES cases tac : context_tactic =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   133
  fn (ctxt, st) => CONTEXT (Proof_Context.update_cases cases ctxt) (tac st);
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   134
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   135
fun CONTEXT_SUBGOAL tac i : context_tactic =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   136
  fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   137
    (case try Logic.nth_prem (i, Thm.prop_of st) of
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   138
      SOME goal => tac (goal, i) (ctxt, st)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   139
    | NONE => Seq.empty);
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   140
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   141
fun (tac1 CONTEXT_THEN_ALL_NEW tac2) i : context_tactic =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   142
  fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   143
    (ctxt, st) |> tac1 i |> Seq.maps_results (fn (ctxt', st') =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   144
      CONTEXT ctxt' ((Seq.INTERVAL tac2 i (i + Thm.nprems_of st' - Thm.nprems_of st)) st'));
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   145
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   146
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   147
(* type method *)
11731
1a0c1ef86518 added trace_rules, trace;
wenzelm
parents: 10907
diff changeset
   148
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   149
type method = thm list -> context_tactic;
11731
1a0c1ef86518 added trace_rules, trace;
wenzelm
parents: 10907
diff changeset
   150
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   151
fun CONTEXT_METHOD tac : method =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   152
  fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac) #> Seq.maps_results (tac facts);
59953
3d207f8f40dd tuned signature;
wenzelm
parents: 59917
diff changeset
   153
3d207f8f40dd tuned signature;
wenzelm
parents: 59917
diff changeset
   154
fun METHOD tac : method =
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   155
  fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac THEN tac facts);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   156
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   157
val fail = METHOD (K no_tac);
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   158
val succeed = METHOD (K all_tac);
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   159
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   160
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   161
(* insert facts *)
7419
wenzelm
parents: 7367
diff changeset
   162
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   163
fun insert_tac _ [] _ = all_tac
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   164
  | insert_tac ctxt facts i =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   165
      EVERY (map (fn r => resolve_tac ctxt [Drule.forall_intr_vars r COMP_INCR revcut_rl] i) facts);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   166
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   167
fun insert thms =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   168
  CONTEXT_METHOD (fn _ => fn (ctxt, st) => st |> ALLGOALS (insert_tac ctxt thms) |> CONTEXT ctxt);
6981
eaade7e398a7 export assumption_tac;
wenzelm
parents: 6951
diff changeset
   169
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   170
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   171
fun SIMPLE_METHOD tac =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   172
  CONTEXT_METHOD (fn facts => fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   173
    st |> (ALLGOALS (insert_tac ctxt facts) THEN tac) |> CONTEXT ctxt);
7419
wenzelm
parents: 7367
diff changeset
   174
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   175
fun SIMPLE_METHOD'' quant tac =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   176
  CONTEXT_METHOD (fn facts => fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   177
    st |> quant (insert_tac ctxt facts THEN' tac) |> CONTEXT ctxt);
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   178
21592
8831206d7f41 renamed SIMPLE_METHOD' to SIMPLE_METHOD'';
wenzelm
parents: 21579
diff changeset
   179
val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL;
9706
8e48a19fc81e removed METHOD0;
wenzelm
parents: 9653
diff changeset
   180
8e48a19fc81e removed METHOD0;
wenzelm
parents: 9653
diff changeset
   181
60578
c708dafe2220 added method "goals" for proper subgoal cases;
wenzelm
parents: 60554
diff changeset
   182
(* goals as cases *)
c708dafe2220 added method "goals" for proper subgoal cases;
wenzelm
parents: 60554
diff changeset
   183
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   184
fun goal_cases_tac case_names : context_tactic =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   185
  fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   186
    let
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   187
      val cases =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   188
        (if null case_names then map string_of_int (1 upto Thm.nprems_of st) else case_names)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   189
        |> map (rpair [] o rpair [])
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   190
        |> Rule_Cases.make_common ctxt (Thm.prop_of (Rule_Cases.internalize_params st));
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   191
    in CONTEXT_CASES cases all_tac (ctxt, st) end;
60578
c708dafe2220 added method "goals" for proper subgoal cases;
wenzelm
parents: 60554
diff changeset
   192
c708dafe2220 added method "goals" for proper subgoal cases;
wenzelm
parents: 60554
diff changeset
   193
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   194
(* cheating *)
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   195
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   196
fun cheating int = CONTEXT_METHOD (fn _ => fn (ctxt, st) =>
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51584
diff changeset
   197
  if int orelse Config.get ctxt quick_and_dirty then
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   198
    CONTEXT ctxt (ALLGOALS (Skip_Proof.cheat_tac ctxt) st)
51552
c713c9505f68 clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents: 51383
diff changeset
   199
  else error "Cheating requires quick_and_dirty mode!");
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   200
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   201
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   202
(* unfold intro/elim rules *)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   203
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   204
fun intro ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (match_tac ctxt ths));
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   205
fun elim ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt ths));
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   206
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   207
12384
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   208
(* unfold/fold definitions *)
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   209
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 33522
diff changeset
   210
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
   211
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
   212
12384
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   213
12829
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   214
(* atomize rule statements *)
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   215
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   216
fun atomize false ctxt =
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   217
      SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac ctxt)
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   218
  | atomize true ctxt =
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   219
      CONTEXT_TACTIC o K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt));
12829
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   220
c92128238f85 full_atomize;
wenzelm
parents: 12399
diff changeset
   221
18039
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   222
(* this -- resolve facts directly *)
12384
86e383f6bfea simple version of 'intro' and 'elim' method;
wenzelm
parents: 12359
diff changeset
   223
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   224
fun this ctxt = METHOD (EVERY o map (HEADGOAL o resolve_tac ctxt o single));
9484
3bda55143260 (un)fold: CHANGED;
wenzelm
parents: 9418
diff changeset
   225
3bda55143260 (un)fold: CHANGED;
wenzelm
parents: 9418
diff changeset
   226
18039
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   227
(* fact -- composition by facts from context *)
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   228
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   229
fun fact [] ctxt = SIMPLE_METHOD' (Proof_Context.some_fact_tac ctxt)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   230
  | fact rules ctxt = SIMPLE_METHOD' (Proof_Context.fact_tac ctxt rules);
18039
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   231
500b7ed7b2bd added fact method;
wenzelm
parents: 17857
diff changeset
   232
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   233
(* assumption *)
7419
wenzelm
parents: 7367
diff changeset
   234
wenzelm
parents: 7367
diff changeset
   235
local
wenzelm
parents: 7367
diff changeset
   236
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   237
fun cond_rtac ctxt cond rule = SUBGOAL (fn (prop, i) =>
19778
f0a318495ca4 assm_tac: try rule termI;
wenzelm
parents: 19482
diff changeset
   238
  if cond (Logic.strip_assums_concl prop)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   239
  then resolve_tac ctxt [rule] i else no_tac);
7419
wenzelm
parents: 7367
diff changeset
   240
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29301
diff changeset
   241
in
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29301
diff changeset
   242
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
   243
fun assm_tac ctxt =
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58957
diff changeset
   244
  assume_tac ctxt APPEND'
23349
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   245
  Goal.assume_rule_tac ctxt APPEND'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   246
  cond_rtac ctxt (can Logic.dest_equals) Drule.reflexive_thm APPEND'
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   247
  cond_rtac ctxt (can Logic.dest_term) Drule.termI;
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   248
49846
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   249
fun all_assm_tac ctxt =
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   250
  let
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   251
    fun tac i st =
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   252
      if i > Thm.nprems_of st then all_tac st
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   253
      else ((assm_tac ctxt i THEN tac i) ORELSE tac (i + 1)) st;
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   254
  in tac 1 end;
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
   255
23349
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   256
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
   257
  (fn [] => assm_tac ctxt
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   258
    | [fact] => solve_tac ctxt [fact]
23349
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   259
    | _ => K no_tac));
23a8345f89f5 Basic text: include position;
wenzelm
parents: 23178
diff changeset
   260
49846
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   261
fun finish immed ctxt =
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58837
diff changeset
   262
  METHOD (K ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac ctxt));
7419
wenzelm
parents: 7367
diff changeset
   263
wenzelm
parents: 7367
diff changeset
   264
end;
wenzelm
parents: 7367
diff changeset
   265
wenzelm
parents: 7367
diff changeset
   266
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   267
(* rule etc. -- single-step refinements *)
12347
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   268
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   269
val rule_trace = Attrib.setup_config_bool \<^binding>\<open>rule_trace\<close> (fn _ => false);
12347
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   270
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   271
fun trace ctxt rules =
41379
b31d7a1cd08f configuration option "rule_trace";
wenzelm
parents: 39507
diff changeset
   272
  if Config.get ctxt rule_trace andalso not (null rules) then
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61166
diff changeset
   273
    Pretty.big_list "rules:" (map (Thm.pretty_thm_item ctxt) rules)
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21879
diff changeset
   274
    |> Pretty.string_of |> tracing
279b129498b6 removed conditional combinator;
wenzelm
parents: 21879
diff changeset
   275
  else ();
12347
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   276
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   277
local
6ee66b76d813 added "rules" method;
wenzelm
parents: 12324
diff changeset
   278
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   279
fun gen_rule_tac tac ctxt rules facts =
18841
edecd40194c1 method (un)folded: option '(raw)';
wenzelm
parents: 18824
diff changeset
   280
  (fn i => fn st =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   281
    if null facts then tac ctxt rules i st
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58837
diff changeset
   282
    else
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   283
      Seq.maps (fn rule => (tac ctxt o single) rule i st)
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58837
diff changeset
   284
        (Drule.multi_resolves (SOME ctxt) facts rules))
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   285
  THEN_ALL_NEW Goal.norm_hhf_tac ctxt;
7130
a17f7b5ac40f added erule;
wenzelm
parents: 6981
diff changeset
   286
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   287
fun gen_arule_tac tac ctxt j rules facts =
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58957
diff changeset
   288
  EVERY' (gen_rule_tac tac ctxt rules facts :: replicate j (assume_tac ctxt));
10744
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   289
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   290
fun gen_some_rule_tac tac ctxt arg_rules facts = SUBGOAL (fn (goal, i) =>
11785
3087d6f19adc intro! and elim! rules;
wenzelm
parents: 11765
diff changeset
   291
  let
3087d6f19adc intro! and elim! rules;
wenzelm
parents: 11765
diff changeset
   292
    val rules =
3087d6f19adc intro! and elim! rules;
wenzelm
parents: 11765
diff changeset
   293
      if not (null arg_rules) then arg_rules
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 60618
diff changeset
   294
      else flat (Context_Rules.find_rules ctxt false facts goal);
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   295
  in trace ctxt rules; tac ctxt rules facts i end);
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10034
diff changeset
   296
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   297
fun meth tac x y = METHOD (HEADGOAL o tac x y);
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   298
fun meth' tac x y z = METHOD (HEADGOAL o tac x y z);
8220
e04928747b18 [df]rule methods;
wenzelm
parents: 8205
diff changeset
   299
7419
wenzelm
parents: 7367
diff changeset
   300
in
wenzelm
parents: 7367
diff changeset
   301
52732
b4da1f2ec73f standardized aliases;
wenzelm
parents: 52059
diff changeset
   302
val rule_tac = gen_rule_tac resolve_tac;
10744
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   303
val rule = meth rule_tac;
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   304
val some_rule_tac = gen_some_rule_tac rule_tac;
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   305
val some_rule = meth some_rule_tac;
10744
5d142ca01b8e 'erule' etc.: assm arg;
wenzelm
parents: 10541
diff changeset
   306
52732
b4da1f2ec73f standardized aliases;
wenzelm
parents: 52059
diff changeset
   307
val erule = meth' (gen_arule_tac eresolve_tac);
b4da1f2ec73f standardized aliases;
wenzelm
parents: 52059
diff changeset
   308
val drule = meth' (gen_arule_tac dresolve_tac);
b4da1f2ec73f standardized aliases;
wenzelm
parents: 52059
diff changeset
   309
val frule = meth' (gen_arule_tac forward_tac);
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   310
7419
wenzelm
parents: 7367
diff changeset
   311
end;
wenzelm
parents: 7367
diff changeset
   312
wenzelm
parents: 7367
diff changeset
   313
25270
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   314
(* intros_tac -- pervasive search spanned by intro rules *)
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   315
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   316
fun gen_intros_tac goals ctxt intros facts =
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   317
  goals (insert_tac ctxt facts THEN'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   318
      REPEAT_ALL_NEW (resolve_tac ctxt intros))
25270
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   319
    THEN Tactic.distinct_subgoals_tac;
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   320
36093
0880493627ca Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents: 33522
diff changeset
   321
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
   322
val try_intros_tac = gen_intros_tac TRYALL;
25270
2ed7b34f58e6 generic tactic Method.intros_tac
haftmann
parents: 24116
diff changeset
   323
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
   324
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   325
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   326
(** method syntax **)
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   327
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   328
(* context data *)
8351
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   329
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   330
structure Data = Generic_Data
26472
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   331
(
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   332
  type T =
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   333
   {methods: ((Token.src -> Proof.context -> method) * string) Name_Space.table,
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   334
    ml_tactic: (morphism -> thm list -> tactic) option,
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   335
    facts: thm list option};
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   336
  val empty : T =
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   337
    {methods = Name_Space.empty_table "method", ml_tactic = NONE, facts = NONE};
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   338
  val extend = I;
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   339
  fun merge
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   340
    ({methods = methods1, ml_tactic = ml_tactic1, facts = facts1},
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   341
     {methods = methods2, ml_tactic = ml_tactic2, facts = facts2}) : T =
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   342
    {methods = Name_Space.merge_tables (methods1, methods2),
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   343
     ml_tactic = merge_options (ml_tactic1, ml_tactic2),
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   344
     facts = merge_options (facts1, facts2)};
26472
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   345
);
9afdd61cf528 ml_tactic: non-critical version via proof data and thread data;
wenzelm
parents: 26463
diff changeset
   346
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   347
fun map_data f = Data.map (fn {methods, ml_tactic, facts} =>
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   348
  let val (methods', ml_tactic', facts') = f (methods, ml_tactic, facts)
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   349
  in {methods = methods', ml_tactic = ml_tactic', facts = facts'} end);
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   350
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   351
val get_methods = #methods o Data.get;
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   352
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   353
val ops_methods =
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   354
 {get_data = get_methods,
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   355
  put_data = fn methods => map_data (fn (_, ml_tactic, facts) => (methods, ml_tactic, facts))};
63090
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents: 62969
diff changeset
   356
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   357
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   358
(* ML tactic *)
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   359
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   360
fun set_tactic ml_tactic = map_data (fn (methods, _, facts) => (methods, SOME ml_tactic, facts));
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   361
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   362
fun the_tactic context =
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   363
  (case #ml_tactic (Data.get context) of
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   364
    SOME tac => tac
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   365
  | NONE => raise Fail "Undefined ML tactic");
8351
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   366
58018
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   367
val parse_tactic =
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   368
  Scan.state :|-- (fn context =>
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   369
    Scan.lift (Args.text_declaration (fn source =>
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   370
      let
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   371
        val context' = context |>
69187
d8849cfad60f clarified ML position: proper markup for def/ref scopes (see also 162a4c2e97bc);
wenzelm
parents: 68823
diff changeset
   372
          ML_Context.expression ("tactic", Position.thread_data ())
d8849cfad60f clarified ML position: proper markup for def/ref scopes (see also 162a4c2e97bc);
wenzelm
parents: 68823
diff changeset
   373
            "Morphism.morphism -> thm list -> tactic"
58991
92b6f4e68c5a more careful ML source positions, for improved PIDE markup;
wenzelm
parents: 58979
diff changeset
   374
            "Method.set_tactic tactic"
59067
dd8ec9138112 tuned signature;
wenzelm
parents: 59064
diff changeset
   375
            (ML_Lex.read "fn morphism: Morphism.morphism => fn facts: thm list =>" @
68823
5e7b1ae10eb8 clarified signature;
wenzelm
parents: 67504
diff changeset
   376
             ML_Lex.read_source source);
58018
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   377
        val tac = the_tactic context';
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   378
      in
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62795
diff changeset
   379
        fn phi => set_tactic (fn _ => Context.setmp_generic_context (SOME context) (tac phi))
58018
beb4b7c0bb30 proper static closure of ML tactic -- data slot is used twice, for ML compiler and transformed declaration;
wenzelm
parents: 58016
diff changeset
   380
      end)) >> (fn decl => Morphism.form (the_tactic (Morphism.form decl context))));
8351
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   381
1b8ac0f48233 added simple_args;
wenzelm
parents: 8335
diff changeset
   382
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   383
(* method facts *)
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   384
63251
9a20078425b1 clean facts more uniformly;
wenzelm
parents: 63250
diff changeset
   385
val clean_facts = filter_out Thm.is_dummy;
9a20078425b1 clean facts more uniformly;
wenzelm
parents: 63250
diff changeset
   386
63256
wenzelm
parents: 63252
diff changeset
   387
fun set_facts facts =
wenzelm
parents: 63252
diff changeset
   388
  (Context.proof_map o map_data)
wenzelm
parents: 63252
diff changeset
   389
    (fn (methods, ml_tactic, _) => (methods, ml_tactic, SOME (clean_facts facts)));
wenzelm
parents: 63252
diff changeset
   390
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   391
val get_facts_generic = these o #facts o Data.get;
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   392
val get_facts = get_facts_generic o Context.Proof;
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   393
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   394
val _ =
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   395
  Theory.setup
64556
851ae0e7b09c more symbols;
wenzelm
parents: 63527
diff changeset
   396
    (Global_Theory.add_thms_dynamic (Binding.make ("method_facts", \<^here>), get_facts_generic));
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   397
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   398
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   399
(* method text *)
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   400
55765
ec7ca5388dea markup for method combinators;
wenzelm
parents: 55761
diff changeset
   401
datatype combinator_info = Combinator_Info of {keywords: Position.T list};
ec7ca5388dea markup for method combinators;
wenzelm
parents: 55761
diff changeset
   402
fun combinator_info keywords = Combinator_Info {keywords = keywords};
ec7ca5388dea markup for method combinators;
wenzelm
parents: 55761
diff changeset
   403
val no_combinator_info = combinator_info [];
ec7ca5388dea markup for method combinators;
wenzelm
parents: 55761
diff changeset
   404
59660
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   405
datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int;
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   406
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   407
datatype text =
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
   408
  Source of Token.src |
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   409
  Basic of Proof.context -> method |
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   410
  Combinator of combinator_info * combinator * text list;
55765
ec7ca5388dea markup for method combinators;
wenzelm
parents: 55761
diff changeset
   411
58006
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   412
fun map_source f (Source src) = Source (f src)
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   413
  | map_source _ (Basic meth) = Basic meth
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   414
  | map_source f (Combinator (info, comb, txts)) = Combinator (info, comb, map (map_source f) txts);
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   415
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54742
diff changeset
   416
fun primitive_text r = Basic (SIMPLE_METHOD o PRIMITIVE o r);
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   417
val succeed_text = Basic (K succeed);
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   418
val standard_text = Source (Token.make_src ("standard", Position.none) []);
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59067
diff changeset
   419
val this_text = Basic this;
32193
c314b4836031 basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
wenzelm
parents: 32091
diff changeset
   420
val done_text = Basic (K (SIMPLE_METHOD all_tac));
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   421
fun sorry_text int = Basic (fn _ => cheating int);
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   422
49846
8fae089f5a0c refined Proof.the_finished_goal with more informative error;
wenzelm
parents: 48992
diff changeset
   423
fun finish_text (NONE, immed) = Basic (finish immed)
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   424
  | finish_text (SOME txt, immed) =
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   425
      Combinator (no_combinator_info, Then, [txt, Basic (finish immed)]);
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   426
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   427
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   428
(* method definitions *)
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   429
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59914
diff changeset
   430
fun print_methods verbose ctxt =
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   431
  let
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   432
    val meths = get_methods (Context.Proof ctxt);
50301
56b4c9afd7be print formal entities with markup;
wenzelm
parents: 49889
diff changeset
   433
    fun prt_meth (name, (_, "")) = Pretty.mark_str name
42813
6c841fa92fa2 optional description for 'attribute_setup' and 'method_setup';
wenzelm
parents: 42616
diff changeset
   434
      | prt_meth (name, (_, comment)) =
50301
56b4c9afd7be print formal entities with markup;
wenzelm
parents: 49889
diff changeset
   435
          Pretty.block
56b4c9afd7be print formal entities with markup;
wenzelm
parents: 49889
diff changeset
   436
            (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment);
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   437
  in
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59914
diff changeset
   438
    [Pretty.big_list "methods:" (map prt_meth (Name_Space.markup_table verbose ctxt meths))]
56334
6b3739fee456 some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents: 56278
diff changeset
   439
    |> Pretty.writeln_chunks
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22586
diff changeset
   440
  end;
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7601
diff changeset
   441
57935
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   442
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   443
(* define *)
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   444
63090
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents: 62969
diff changeset
   445
fun define_global binding meth comment =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents: 62969
diff changeset
   446
  Entity.define_global ops_methods binding (meth, comment);
57935
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   447
57941
57200bdc2aa7 localized command 'method_setup' and 'attribute_setup';
wenzelm
parents: 57938
diff changeset
   448
fun define binding meth comment =
63090
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents: 62969
diff changeset
   449
  Entity.define ops_methods binding (meth, comment);
31304
00a9c674cf40 eliminated old Method.add_method(s);
wenzelm
parents: 31303
diff changeset
   450
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55917
diff changeset
   451
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55917
diff changeset
   452
(* check *)
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55917
diff changeset
   453
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   454
fun check_name ctxt =
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   455
  let val context = Context.Proof ctxt
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   456
  in #1 o Name_Space.check context (get_methods context) end;
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   457
59907
6c0f62490699 tuned signature;
wenzelm
parents: 59666
diff changeset
   458
fun check_src ctxt =
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   459
  #1 o Token.check_src ctxt (get_methods o Context.Proof);
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55917
diff changeset
   460
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   461
fun check_text ctxt (Source src) = Source (check_src ctxt src)
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   462
  | check_text _ (Basic m) = Basic m
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   463
  | check_text ctxt (Combinator (x, y, body)) = Combinator (x, y, map (check_text ctxt) body);
59914
d1ddcd8df4e4 proper treatment of internal method name as already checked Token.src;
wenzelm
parents: 59909
diff changeset
   464
61917
35ec3757d3c1 check and report source at most once, notably in body of "match" method;
wenzelm
parents: 61843
diff changeset
   465
fun checked_text (Source src) = Token.checked_src src
35ec3757d3c1 check and report source at most once, notably in body of "match" method;
wenzelm
parents: 61843
diff changeset
   466
  | checked_text (Basic _) = true
35ec3757d3c1 check and report source at most once, notably in body of "match" method;
wenzelm
parents: 61843
diff changeset
   467
  | checked_text (Combinator (_, _, body)) = forall checked_text body;
35ec3757d3c1 check and report source at most once, notably in body of "match" method;
wenzelm
parents: 61843
diff changeset
   468
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55917
diff changeset
   469
30512
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
   470
(* method setup *)
17b2aad869fa added simplified setup;
wenzelm
parents: 30508
diff changeset
   471
57935
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   472
fun method_syntax scan src ctxt : method =
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
   473
  let val (m, ctxt') = Token.syntax scan src ctxt in m ctxt' end;
57935
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   474
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   475
fun setup binding scan comment = define_global binding (method_syntax scan) comment #> snd;
c578f3a37a67 localized method definitions (see also f14c1248d064);
wenzelm
parents: 57863
diff changeset
   476
fun local_setup binding scan comment = define binding (method_syntax scan) comment #> snd;
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   477
69187
d8849cfad60f clarified ML position: proper markup for def/ref scopes (see also 162a4c2e97bc);
wenzelm
parents: 68823
diff changeset
   478
fun method_setup (name, pos) source comment =
68823
5e7b1ae10eb8 clarified signature;
wenzelm
parents: 67504
diff changeset
   479
  ML_Lex.read_source source
69187
d8849cfad60f clarified ML position: proper markup for def/ref scopes (see also 162a4c2e97bc);
wenzelm
parents: 68823
diff changeset
   480
  |> ML_Context.expression ("parser", pos)
58991
92b6f4e68c5a more careful ML source positions, for improved PIDE markup;
wenzelm
parents: 58979
diff changeset
   481
    "(Proof.context -> Proof.method) context_parser"
69187
d8849cfad60f clarified ML position: proper markup for def/ref scopes (see also 162a4c2e97bc);
wenzelm
parents: 68823
diff changeset
   482
    ("Context.map_proof (Method.local_setup " ^
d8849cfad60f clarified ML position: proper markup for def/ref scopes (see also 162a4c2e97bc);
wenzelm
parents: 68823
diff changeset
   483
      ML_Syntax.atomic (ML_Syntax.make_binding (name, pos)) ^
58979
162a4c2e97bc more careful ML source positions, for improved PIDE markup;
wenzelm
parents: 58978
diff changeset
   484
      " parser " ^ ML_Syntax.print_string comment ^ ")")
57941
57200bdc2aa7 localized command 'method_setup' and 'attribute_setup';
wenzelm
parents: 57938
diff changeset
   485
  |> Context.proof_map;
17356
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   486
09afdf37cdb3 added cheating, sorry_text (from skip_proofs.ML);
wenzelm
parents: 17314
diff changeset
   487
58006
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   488
(* prepare methods *)
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   489
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   490
fun method ctxt =
58016
28e5ccf4a27f more uniform data slot;
wenzelm
parents: 58011
diff changeset
   491
  let val table = get_methods (Context.Proof ctxt)
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 58007
diff changeset
   492
  in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end;
58006
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   493
59909
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
   494
fun method_closure ctxt src =
58006
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   495
  let
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   496
    val src' = map Token.init_assignable src;
59909
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
   497
    val ctxt' = Context_Position.not_really ctxt;
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   498
    val _ = Seq.pull (method ctxt' src' ctxt' [] (ctxt', Goal.protect 0 Drule.dummy_thm));
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   499
  in map Token.closure src' end;
58006
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   500
64556
851ae0e7b09c more symbols;
wenzelm
parents: 63527
diff changeset
   501
val closure = Config.bool (Config.declare ("Method.closure", \<^here>) (K (Config.Bool true)));
59909
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
   502
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
   503
fun method_cmd ctxt =
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
   504
  check_src ctxt #>
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
   505
  Config.get ctxt closure ? method_closure ctxt #>
fbf4d5ad500d clarified method_closure;
wenzelm
parents: 59907
diff changeset
   506
  method ctxt;
58006
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   507
3072ff7ea472 tuned signature;
wenzelm
parents: 58005
diff changeset
   508
60609
15620ae824c0 clarified static phase;
wenzelm
parents: 60578
diff changeset
   509
(* static vs. runtime state *)
60553
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   510
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   511
fun detect_closure_state st =
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   512
  (case try Logic.dest_term (Thm.concl_of (perhaps (try Goal.conclude) st)) of
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   513
    NONE => false
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   514
  | SOME t => Term.is_dummy_pattern t);
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   515
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   516
fun STATIC test : context_tactic =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   517
  fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   518
    if detect_closure_state st then (test (); Seq.single (Seq.Result (ctxt, st))) else Seq.empty;
60609
15620ae824c0 clarified static phase;
wenzelm
parents: 60578
diff changeset
   519
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   520
fun RUNTIME (tac: context_tactic) (ctxt, st) =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   521
  if detect_closure_state st then Seq.empty else tac (ctxt, st);
60553
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   522
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   523
fun sleep t = RUNTIME (fn ctxt_st => (OS.Process.sleep t; Seq.single (Seq.Result ctxt_st)));
60554
c0e1c121c7c0 added method "sleep";
wenzelm
parents: 60553
diff changeset
   524
60553
86fc6652c4df tuned signature;
wenzelm
parents: 60212
diff changeset
   525
58003
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   526
(* evaluate method text *)
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   527
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   528
local
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   529
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   530
val op THEN = Seq.THEN;
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   531
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   532
fun BYPASS_CONTEXT (tac: tactic) =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   533
  fn result =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   534
    (case result of
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   535
      Seq.Error _ => Seq.single result
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   536
    | Seq.Result (ctxt, st) => tac st |> CONTEXT ctxt);
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   537
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   538
val preparation = BYPASS_CONTEXT (ALLGOALS Goal.conjunction_tac);
59660
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   539
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   540
fun RESTRICT_GOAL i n method =
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   541
  BYPASS_CONTEXT (PRIMITIVE (Goal.restrict i n)) THEN
59660
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   542
  method THEN
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   543
  BYPASS_CONTEXT (PRIMITIVE (Goal.unrestrict i));
59660
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   544
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   545
fun SELECT_GOAL method i = RESTRICT_GOAL i 1 method;
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   546
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   547
fun (method1 THEN_ALL_NEW method2) i (result : context_state Seq.result) =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   548
  (case result of
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   549
    Seq.Error _ => Seq.single result
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   550
  | Seq.Result (_, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   551
      result |> method1 i
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   552
      |> Seq.maps (fn result' =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   553
          (case result' of
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   554
            Seq.Error _ => Seq.single result'
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   555
          | Seq.Result (_, st') =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   556
              result' |> Seq.INTERVAL method2 i (i + Thm.nprems_of st' - Thm.nprems_of st))))
58003
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   557
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   558
fun COMBINATOR1 comb [meth] = comb meth
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   559
  | COMBINATOR1 _ _ = raise Fail "Method combinator requires exactly one argument";
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   560
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   561
fun combinator Then = Seq.EVERY
59660
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   562
  | combinator Then_All_New =
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   563
      (fn [] => Seq.single
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   564
        | methods =>
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   565
            preparation THEN (foldl1 (op THEN_ALL_NEW) (map SELECT_GOAL methods) 1))
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   566
  | combinator Orelse = Seq.FIRST
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   567
  | combinator Try = COMBINATOR1 Seq.TRY
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   568
  | combinator Repeat1 = COMBINATOR1 Seq.REPEAT1
59660
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   569
  | combinator (Select_Goals n) =
49e498cedd02 support structural composition (THEN_ALL_NEW) for proof methods;
wenzelm
parents: 59498
diff changeset
   570
      COMBINATOR1 (fn method => preparation THEN RESTRICT_GOAL 1 n method);
58003
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   571
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   572
in
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   573
63257
94d1f820130f provide dynamic facts in static context, to allow use of method_facts during static closure;
wenzelm
parents: 63256
diff changeset
   574
fun evaluate text ctxt0 facts =
58003
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   575
  let
63257
94d1f820130f provide dynamic facts in static context, to allow use of method_facts during static closure;
wenzelm
parents: 63256
diff changeset
   576
    val ctxt = set_facts facts ctxt0;
94d1f820130f provide dynamic facts in static context, to allow use of method_facts during static closure;
wenzelm
parents: 63256
diff changeset
   577
    fun eval0 m = Seq.single #> Seq.maps_results (m facts);
94d1f820130f provide dynamic facts in static context, to allow use of method_facts during static closure;
wenzelm
parents: 63256
diff changeset
   578
    fun eval (Basic m) = eval0 (m ctxt)
94d1f820130f provide dynamic facts in static context, to allow use of method_facts during static closure;
wenzelm
parents: 63256
diff changeset
   579
      | eval (Source src) = eval0 (method_cmd ctxt src ctxt)
63250
96cfb07c60d4 expode method_facts via dynamic method context;
wenzelm
parents: 63090
diff changeset
   580
      | eval (Combinator (_, c, txts)) = combinator c (map eval txts);
63257
94d1f820130f provide dynamic facts in static context, to allow use of method_facts during static closure;
wenzelm
parents: 63256
diff changeset
   581
  in eval text o Seq.Result end;
58003
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   582
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   583
end;
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   584
63527
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   585
fun evaluate_runtime text ctxt =
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   586
  let
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   587
    val text' =
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   588
      text |> (map_source o map o Token.map_facts)
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   589
        (fn SOME name =>
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   590
              (case Proof_Context.lookup_fact ctxt name of
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   591
                SOME {dynamic = true, thms} => K thms
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   592
              | _ => I)
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   593
          | NONE => I);
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   594
    val ctxt' = Config.put closure false ctxt;
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   595
  in fn facts => RUNTIME (fn st => evaluate text' ctxt' facts st) end;
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   596
58003
250ecd2502ad clarifed Method.evaluate: turn text into semantic method (like Basic);
wenzelm
parents: 58002
diff changeset
   597
5884
113badd4dae5 several args parsers;
wenzelm
parents: 5824
diff changeset
   598
17110
4c5622d7bdbe moved before proof.ML;
wenzelm
parents: 16876
diff changeset
   599
(** concrete syntax **)
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   600
58048
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   601
(* type modifier *)
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   602
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   603
type modifier =
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   604
  {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T};
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   605
58048
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   606
fun modifier attribute pos : modifier = {init = I, attribute = attribute, pos = pos};
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   607
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   608
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   609
(* sections *)
7268
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   610
59982
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   611
val old_section_parser =
64556
851ae0e7b09c more symbols;
wenzelm
parents: 63527
diff changeset
   612
  Config.bool (Config.declare ("Method.old_section_parser", \<^here>) (K (Config.Bool false)));
59982
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   613
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   614
local
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   615
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   616
fun thms ss =
61476
1884c40f1539 tuned signature;
wenzelm
parents: 61268
diff changeset
   617
  Scan.repeats (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm);
59982
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   618
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   619
fun app {init, attribute, pos = _} ths context =
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   620
  fold_map (Thm.apply_attribute attribute) ths (Context.map_proof init context);
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   621
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   622
fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|--
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   623
  (fn (m, ths) => Scan.succeed (swap (app m ths context))));
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   624
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   625
in
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   626
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   627
fun old_sections ss = Scan.repeat (section ss) >> K ();
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   628
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   629
end;
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   630
7268
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   631
local
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   632
58048
aa6296d09e0e more explicit Method.modifier with reported position;
wenzelm
parents: 58034
diff changeset
   633
fun sect (modifier : modifier parser) = Scan.depend (fn context =>
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62889
diff changeset
   634
  Scan.ahead Parse.not_eof -- Scan.trace modifier -- Scan.repeat (Scan.unless modifier Parse.thm)
60211
c0f686b39ebb modifier markup for all parsed tokens;
wenzelm
parents: 59982
diff changeset
   635
    >> (fn ((tok0, ({init, attribute, pos}, modifier_toks)), xthms) =>
58029
2137e60b6f6d clarified Method.section: explicit declaration with static closure;
wenzelm
parents: 58026
diff changeset
   636
      let
58068
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   637
        val decl =
60211
c0f686b39ebb modifier markup for all parsed tokens;
wenzelm
parents: 59982
diff changeset
   638
          (case Token.get_value tok0 of
58068
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   639
            SOME (Token.Declaration decl) => decl
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   640
          | _ =>
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   641
              let
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   642
                val ctxt = Context.proof_of context;
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   643
                val prep_att = Attrib.check_src ctxt #> map (Token.assign NONE);
58068
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   644
                val thms =
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   645
                  map (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs)) xthms;
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   646
                val facts =
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   647
                  Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)]
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   648
                  |> map (fn (_, bs) => ((Binding.empty, [Attrib.internal (K attribute)]), bs));
60211
c0f686b39ebb modifier markup for all parsed tokens;
wenzelm
parents: 59982
diff changeset
   649
58068
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   650
                fun decl phi =
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   651
                  Context.mapping I init #>
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   652
                  Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd;
60211
c0f686b39ebb modifier markup for all parsed tokens;
wenzelm
parents: 59982
diff changeset
   653
60212
4e8d8baa491c clarified markup range;
wenzelm
parents: 60211
diff changeset
   654
                val modifier_report =
62795
063d2f23cdf6 removed redundant Position.set_range -- already done in Position.range;
wenzelm
parents: 61917
diff changeset
   655
                  (#1 (Token.range_of modifier_toks),
60212
4e8d8baa491c clarified markup range;
wenzelm
parents: 60211
diff changeset
   656
                    Markup.properties (Position.def_properties_of pos)
4e8d8baa491c clarified markup range;
wenzelm
parents: 60211
diff changeset
   657
                      (Markup.entity Markup.method_modifierN ""));
60211
c0f686b39ebb modifier markup for all parsed tokens;
wenzelm
parents: 59982
diff changeset
   658
                val _ =
60212
4e8d8baa491c clarified markup range;
wenzelm
parents: 60211
diff changeset
   659
                  Context_Position.reports ctxt (modifier_report :: Token.reports_of_value tok0);
60211
c0f686b39ebb modifier markup for all parsed tokens;
wenzelm
parents: 59982
diff changeset
   660
                val _ = Token.assign (SOME (Token.Declaration decl)) tok0;
58068
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   661
              in decl end);
d6f29bf9b783 intern xthm only once;
wenzelm
parents: 58048
diff changeset
   662
      in (Morphism.form decl context, decl) end));
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   663
30540
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
   664
in
5e2d9604a3d3 export section, sections;
wenzelm
parents: 30515
diff changeset
   665
59982
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   666
fun sections ss =
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   667
  Args.context :|-- (fn ctxt =>
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   668
    if Config.get ctxt old_section_parser then old_sections ss
f402fd001429 option for old section parser (before 2137e60b6f6d) for the sake of Eisbach;
wenzelm
parents: 59981
diff changeset
   669
    else Scan.repeat (sect (Scan.first ss)) >> K ());
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   670
7268
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   671
end;
315655dc361b sectioned_args etc.: more general modifier;
wenzelm
parents: 7130
diff changeset
   672
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   673
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   674
(* extra rule methods *)
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   675
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   676
fun xrule_meth meth =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36096
diff changeset
   677
  Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >>
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   678
  (fn (n, ths) => fn ctxt => meth ctxt n ths);
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   679
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   680
55761
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   681
(* text range *)
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   682
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   683
type text_range = text * Position.range;
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   684
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   685
fun text NONE = NONE
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   686
  | text (SOME (txt, _)) = SOME txt;
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   687
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   688
fun position NONE = Position.none
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   689
  | position (SOME (_, (pos, _))) = pos;
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   690
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   691
55795
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   692
(* reports *)
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   693
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   694
local
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   695
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   696
fun keyword_positions (Source _) = []
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   697
  | keyword_positions (Basic _) = []
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   698
  | keyword_positions (Combinator (Combinator_Info {keywords}, _, texts)) =
c28e6bc6635d more compact datatypes;
wenzelm
parents: 58004
diff changeset
   699
      keywords @ maps keyword_positions texts;
55795
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   700
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   701
in
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   702
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   703
fun reports_of ((text, (pos, _)): text_range) =
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   704
  (pos, Markup.language_method) ::
55917
5438ed05e1c9 special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
wenzelm
parents: 55828
diff changeset
   705
    maps (fn p => map (pair p) (Markup.keyword3 :: Completion.suppress_abbrevs ""))
5438ed05e1c9 special treatment of method combinators like Args.$$$ keywords, although parsed via Parse.$$$;
wenzelm
parents: 55828
diff changeset
   706
      (keyword_positions text);
55795
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   707
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   708
val report = Position.reports o reports_of;
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   709
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   710
end;
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   711
2d4cf0005a02 more explicit method reports;
wenzelm
parents: 55765
diff changeset
   712
59666
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   713
(* parser *)
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   714
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   715
local
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   716
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   717
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
   718
  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
   719
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   720
in
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   721
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   722
fun parser pri =
59666
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   723
  let
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62889
diff changeset
   724
    val meth_name = Parse.token Parse.name;
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   725
59666
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   726
    fun meth5 x =
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   727
     (meth_name >> (Source o single) ||
59666
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   728
      Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok =>
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   729
        Source (Token.make_src ("cartouche", Position.none) [tok])) ||
59666
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   730
      Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   731
    and meth4 x =
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   732
     (meth5 -- Parse.position (Parse.$$$ "?")
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   733
        >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Try, [m])) ||
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   734
      meth5 -- Parse.position (Parse.$$$ "+")
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   735
        >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Repeat1, [m])) ||
59914
d1ddcd8df4e4 proper treatment of internal method name as already checked Token.src;
wenzelm
parents: 59909
diff changeset
   736
      meth5 -- (Parse.position (Parse.$$$ "[") --
d1ddcd8df4e4 proper treatment of internal method name as already checked Token.src;
wenzelm
parents: 59909
diff changeset
   737
          Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]"))
59666
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   738
        >> (fn (m, (((_, pos1), n), (_, pos2))) =>
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   739
            Combinator (combinator_info [pos1, pos2], Select_Goals n, [m])) ||
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   740
      meth5) x
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   741
    and meth3 x =
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61812
diff changeset
   742
     (meth_name ::: Parse.args1 is_symid_meth >> Source ||
59666
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   743
      meth4) x
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   744
    and meth2 x =
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   745
      (Parse.enum1_positions "," meth3
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   746
        >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then, ms))) x
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   747
    and meth1 x =
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   748
      (Parse.enum1_positions ";" meth2
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   749
        >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then_All_New, ms))) x
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   750
    and meth0 x =
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   751
      (Parse.enum1_positions "|" meth1
0e9f303d1515 clarified Token.check_src: intern at most once;
wenzelm
parents: 59660
diff changeset
   752
        >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Orelse, ms))) x;
59981
0d0f9c66ff3f tuned signature;
wenzelm
parents: 59953
diff changeset
   753
0d0f9c66ff3f tuned signature;
wenzelm
parents: 59953
diff changeset
   754
    val meth =
0d0f9c66ff3f tuned signature;
wenzelm
parents: 59953
diff changeset
   755
      nth [meth0, meth1, meth2, meth3, meth4, meth5] pri
0d0f9c66ff3f tuned signature;
wenzelm
parents: 59953
diff changeset
   756
        handle General.Subscript => raise Fail ("Bad method parser priority " ^ string_of_int pri);
0d0f9c66ff3f tuned signature;
wenzelm
parents: 59953
diff changeset
   757
  in Scan.trace meth >> (fn (m, toks) => (m, Token.range_of toks)) end;
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   758
59981
0d0f9c66ff3f tuned signature;
wenzelm
parents: 59953
diff changeset
   759
val parse = parser 4;
49866
619acbd72664 more proof method text position information;
wenzelm
parents: 49865
diff changeset
   760
55761
213b9811f59f method language markup, e.g. relevant to prevent outer keyword completion;
wenzelm
parents: 55742
diff changeset
   761
end;
27813
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   762
96fbe385a0d0 unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 27751
diff changeset
   763
63527
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   764
(* read method text *)
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   765
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   766
fun read ctxt src =
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   767
  (case Scan.read Token.stopper (Parse.!!! (parser 0 --| Scan.ahead Parse.eof)) src of
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   768
    SOME (text, range) =>
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   769
      if checked_text text then text
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   770
      else (report (text, range); check_text ctxt text)
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   771
  | NONE => error ("Failed to parse method" ^ Position.here (#1 (Token.range_of src))));
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   772
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   773
fun read_closure ctxt src0 =
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   774
  let
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   775
    val src1 = map Token.init_assignable src0;
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   776
    val text = read ctxt src1 |> map_source (method_closure ctxt);
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   777
    val src2 = map Token.closure src1;
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   778
  in (text, src2) end;
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   779
67504
310114bec0d7 clarified signature;
wenzelm
parents: 67502
diff changeset
   780
fun read_closure_input ctxt input =
310114bec0d7 clarified signature;
wenzelm
parents: 67502
diff changeset
   781
  let val syms = Input.source_explode input in
310114bec0d7 clarified signature;
wenzelm
parents: 67502
diff changeset
   782
    (case Token.read_body (Thy_Header.get_keywords' ctxt) (Scan.many Token.not_eof) syms of
310114bec0d7 clarified signature;
wenzelm
parents: 67502
diff changeset
   783
      SOME res => read_closure ctxt res
310114bec0d7 clarified signature;
wenzelm
parents: 67502
diff changeset
   784
    | NONE => error ("Bad input" ^ Position.here (Input.pos_of input)))
310114bec0d7 clarified signature;
wenzelm
parents: 67502
diff changeset
   785
  end;
63527
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   786
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   787
val text_closure =
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   788
  Args.context -- Scan.lift (Parse.token Parse.text) >> (fn (ctxt, tok) =>
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   789
    (case Token.get_value tok of
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   790
      SOME (Token.Source src) => read ctxt src
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   791
    | _ =>
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   792
        let
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   793
          val (text, src) = read_closure_input ctxt (Token.input_of tok);
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   794
          val _ = Token.assign (SOME (Token.Source src)) tok;
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   795
        in text end));
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   796
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   797
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18640
diff changeset
   798
(* theory setup *)
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   799
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 53168
diff changeset
   800
val _ = Theory.setup
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   801
 (setup \<^binding>\<open>fail\<close> (Scan.succeed (K fail)) "force failure" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   802
  setup \<^binding>\<open>succeed\<close> (Scan.succeed (K succeed)) "succeed" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   803
  setup \<^binding>\<open>sleep\<close> (Scan.lift Parse.real >> (fn s => fn _ => fn _ => sleep (seconds s)))
60554
c0e1c121c7c0 added method "sleep";
wenzelm
parents: 60553
diff changeset
   804
    "succeed after delay (in seconds)" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   805
  setup \<^binding>\<open>-\<close> (Scan.succeed (K (SIMPLE_METHOD all_tac)))
60578
c708dafe2220 added method "goals" for proper subgoal cases;
wenzelm
parents: 60554
diff changeset
   806
    "insert current facts, nothing else" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   807
  setup \<^binding>\<open>goal_cases\<close> (Scan.lift (Scan.repeat Args.name_token) >> (fn names => fn _ =>
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   808
    CONTEXT_METHOD (fn _ => fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   809
      (case drop (Thm.nprems_of st) names of
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   810
        [] => NONE
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   811
      | bad =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   812
          if detect_closure_state st then NONE
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   813
          else
61843
1803599838a6 tuned message;
wenzelm
parents: 61841
diff changeset
   814
            SOME (fn () => ("Excessive case name(s): " ^ commas_quote (map Token.content_of bad) ^
62795
063d2f23cdf6 removed redundant Position.set_range -- already done in Position.range;
wenzelm
parents: 61917
diff changeset
   815
              Position.here (#1 (Token.range_of bad)))))
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   816
      |> (fn SOME msg => Seq.single (Seq.Error msg)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   817
           | NONE => goal_cases_tac (map Token.content_of names) (ctxt, st)))))
61164
2a03ae772c60 method "goals" ignores facts;
wenzelm
parents: 61049
diff changeset
   818
    "bind cases for goals" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   819
  setup \<^binding>\<open>insert\<close> (Attrib.thms >> (K o insert))
60578
c708dafe2220 added method "goals" for proper subgoal cases;
wenzelm
parents: 60554
diff changeset
   820
    "insert theorems, ignoring facts" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   821
  setup \<^binding>\<open>intro\<close> (Attrib.thms >> (fn ths => fn ctxt => intro ctxt ths))
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   822
    "repeatedly apply introduction rules" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   823
  setup \<^binding>\<open>elim\<close> (Attrib.thms >> (fn ths => fn ctxt => elim ctxt ths))
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   824
    "repeatedly apply elimination rules" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   825
  setup \<^binding>\<open>unfold\<close> (Attrib.thms >> unfold_meth) "unfold definitions" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   826
  setup \<^binding>\<open>fold\<close> (Attrib.thms >> fold_meth) "fold definitions" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   827
  setup \<^binding>\<open>atomize\<close> (Scan.lift (Args.mode "full") >> atomize)
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   828
    "present local premises as object-level statements" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   829
  setup \<^binding>\<open>rule\<close> (Attrib.thms >> (fn ths => fn ctxt => some_rule ctxt ths))
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53171
diff changeset
   830
    "apply some intro/elim rule" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   831
  setup \<^binding>\<open>erule\<close> (xrule_meth erule) "apply rule in elimination manner (improper)" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   832
  setup \<^binding>\<open>drule\<close> (xrule_meth drule) "apply rule in destruct manner (improper)" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   833
  setup \<^binding>\<open>frule\<close> (xrule_meth frule) "apply rule in forward manner (improper)" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   834
  setup \<^binding>\<open>this\<close> (Scan.succeed this) "apply current facts as rules" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   835
  setup \<^binding>\<open>fact\<close> (Attrib.thms >> fact) "composition by facts from context" #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   836
  setup \<^binding>\<open>assumption\<close> (Scan.succeed assumption)
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   837
    "proof by assumption, preferring facts" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   838
  setup \<^binding>\<open>rename_tac\<close> (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >>
52732
b4da1f2ec73f standardized aliases;
wenzelm
parents: 52059
diff changeset
   839
    (fn (quant, xs) => K (SIMPLE_METHOD'' quant (rename_tac xs))))
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   840
    "rename parameters of goal" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   841
  setup \<^binding>\<open>rotate_tac\<close> (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >>
52732
b4da1f2ec73f standardized aliases;
wenzelm
parents: 52059
diff changeset
   842
    (fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i))))
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   843
      "rotate assumptions of goal" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   844
  setup \<^binding>\<open>tactic\<close> (parse_tactic >> (K o METHOD))
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30512
diff changeset
   845
    "ML tactic as proof method" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   846
  setup \<^binding>\<open>raw_tactic\<close> (parse_tactic >> (fn tac => fn _ => CONTEXT_TACTIC o tac))
63527
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   847
    "ML tactic as raw proof method" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64556
diff changeset
   848
  setup \<^binding>\<open>use\<close>
63527
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   849
    (Attrib.thms -- (Scan.lift (Parse.$$$ "in") |-- text_closure) >>
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   850
      (fn (thms, text) => fn ctxt => fn _ => evaluate_runtime text ctxt thms))
59eff6e56d81 moved method "use" to Pure;
wenzelm
parents: 63257
diff changeset
   851
    "indicate method facts and context for method expression");
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   852
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   853
16145
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   854
(*final declarations of this structure!*)
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   855
val unfold = unfold_meth;
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   856
val fold = fold_meth;
1bb17485602f renamed cond_extern to extern;
wenzelm
parents: 15973
diff changeset
   857
5824
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   858
end;
91113aa09371 Proof methods.
wenzelm
parents:
diff changeset
   859
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   860
structure Basic_Method: BASIC_METHOD = Method;
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   861
open Basic_Method;
2154e6c8d52d clarified modules;
wenzelm
parents: 61814
diff changeset
   862
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61834
diff changeset
   863
val CONTEXT_METHOD = Method.CONTEXT_METHOD;
30508
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   864
val METHOD = Method.METHOD;
958cc116d03b tuned Method exports: non-pervasive type method (cf. Proof.method), pervasive METHOD combinators;
wenzelm
parents: 30466
diff changeset
   865
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
   866
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
   867
val SIMPLE_METHOD'' = Method.SIMPLE_METHOD'';