src/Doc/Implementation/Isar.thy
author wenzelm
Sat, 05 Jan 2019 17:24:33 +0100
changeset 69597 ff784d5a5bfb
parent 62969 9f394a16c557
child 73764 35d8132633c6
permissions -rw-r--r--
isabelle update -u control_cartouches;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61656
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61572
diff changeset
     1
(*:maxLineLen=78:*)
cfabbc083977 more uniform jEdit properties;
wenzelm
parents: 61572
diff changeset
     2
29755
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 20520
diff changeset
     3
theory Isar
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 20520
diff changeset
     4
imports Base
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 20520
diff changeset
     5
begin
20472
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
     7
chapter \<open>Isar language elements\<close>
29759
bcb79ddf57da removed rudiments of glossary;
wenzelm
parents: 29755
diff changeset
     8
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
     9
text \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    10
  The Isar proof language (see also @{cite \<open>\S2\<close> "isabelle-isar-ref"})
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    11
  consists of three main categories of language elements:
29759
bcb79ddf57da removed rudiments of glossary;
wenzelm
parents: 29755
diff changeset
    12
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    13
  \<^enum> Proof \<^emph>\<open>commands\<close> define the primary language of transactions of the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    14
  underlying Isar/VM interpreter. Typical examples are @{command "fix"},
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    15
  @{command "assume"}, @{command "show"}, @{command "proof"}, and @{command
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    16
  "qed"}.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    17
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    18
  Composing proof commands according to the rules of the Isar/VM leads to
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    19
  expressions of structured proof text, such that both the machine and the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    20
  human reader can give it a meaning as formal reasoning.
20472
wenzelm
parents:
diff changeset
    21
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    22
  \<^enum> Proof \<^emph>\<open>methods\<close> define a secondary language of mixed forward-backward
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    23
  refinement steps involving facts and goals. Typical examples are @{method
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    24
  rule}, @{method unfold}, and @{method simp}.
29759
bcb79ddf57da removed rudiments of glossary;
wenzelm
parents: 29755
diff changeset
    25
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    26
  Methods can occur in certain well-defined parts of the Isar proof language,
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    27
  say as arguments to @{command "proof"}, @{command "qed"}, or @{command
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    28
  "by"}.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    29
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    30
  \<^enum> \<^emph>\<open>Attributes\<close> define a tertiary language of small annotations to theorems
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    31
  being defined or referenced. Attributes can modify both the context and the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    32
  theorem.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    33
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    34
  Typical examples are @{attribute intro} (which affects the context), and
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    35
  @{attribute symmetric} (which affects the theorem).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    36
\<close>
29759
bcb79ddf57da removed rudiments of glossary;
wenzelm
parents: 29755
diff changeset
    37
bcb79ddf57da removed rudiments of glossary;
wenzelm
parents: 29755
diff changeset
    38
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    39
section \<open>Proof commands\<close>
20520
wenzelm
parents: 20472
diff changeset
    40
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    41
text \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    42
  A \<^emph>\<open>proof command\<close> is state transition of the Isar/VM proof interpreter.
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
    43
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    44
  In principle, Isar proof commands could be defined in user-space as well.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    45
  The system is built like that in the first place: one part of the commands
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    46
  are primitive, the other part is defined as derived elements. Adding to the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    47
  genuine structured proof language requires profound understanding of the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    48
  Isar/VM machinery, though, so this is beyond the scope of this manual.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    49
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    50
  What can be done realistically is to define some diagnostic commands that
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    51
  inspect the general state of the Isar/VM, and report some feedback to the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    52
  user. Typically this involves checking of the linguistic \<^emph>\<open>mode\<close> of a proof
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    53
  state, or peeking at the pending goals (if available).
39845
50f42116ebdb more on Proof.theorem;
wenzelm
parents: 39844
diff changeset
    54
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    55
  Another common application is to define a toplevel command that poses a
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    56
  problem to the user as Isar proof state and processes the final result
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    57
  relatively to the context. Thus a proof can be incorporated into the context
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    58
  of some user-space tool, without modifying the Isar proof language itself.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    59
\<close>
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    60
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    61
text %mlref \<open>
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    62
  \begin{mldecls}
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    63
  @{index_ML_type Proof.state} \\
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    64
  @{index_ML Proof.assert_forward: "Proof.state -> Proof.state"} \\
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    65
  @{index_ML Proof.assert_chain: "Proof.state -> Proof.state"} \\
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    66
  @{index_ML Proof.assert_backward: "Proof.state -> Proof.state"} \\
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    67
  @{index_ML Proof.simple_goal: "Proof.state -> {context: Proof.context, goal: thm}"} \\
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    68
  @{index_ML Proof.goal: "Proof.state ->
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    69
  {context: Proof.context, facts: thm list, goal: thm}"} \\
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    70
  @{index_ML Proof.raw_goal: "Proof.state ->
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    71
  {context: Proof.context, facts: thm list, goal: thm}"} \\
39845
50f42116ebdb more on Proof.theorem;
wenzelm
parents: 39844
diff changeset
    72
  @{index_ML Proof.theorem: "Method.text option ->
50f42116ebdb more on Proof.theorem;
wenzelm
parents: 39844
diff changeset
    73
  (thm list list -> Proof.context -> Proof.context) ->
50f42116ebdb more on Proof.theorem;
wenzelm
parents: 39844
diff changeset
    74
  (term * term list) list list -> Proof.context -> Proof.state"} \\
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    75
  \end{mldecls}
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    76
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
    77
  \<^descr> Type \<^ML_type>\<open>Proof.state\<close> represents Isar proof states. This is a
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    78
  block-structured configuration with proof context, linguistic mode, and
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    79
  optional goal. The latter consists of goal context, goal facts
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    80
  (``\<open>using\<close>''), and tactical goal state (see \secref{sec:tactical-goals}).
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    81
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    82
  The general idea is that the facts shall contribute to the refinement of
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    83
  some parts of the tactical goal --- how exactly is defined by the proof
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    84
  method that is applied in that situation.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    85
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
    86
  \<^descr> \<^ML>\<open>Proof.assert_forward\<close>, \<^ML>\<open>Proof.assert_chain\<close>, \<^ML>\<open>Proof.assert_backward\<close> are partial identity functions that fail unless a
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    87
  certain linguistic mode is active, namely ``\<open>proof(state)\<close>'',
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    88
  ``\<open>proof(chain)\<close>'', ``\<open>proof(prove)\<close>'', respectively (using the terminology
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    89
  of @{cite "isabelle-isar-ref"}).
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    90
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    91
  It is advisable study the implementations of existing proof commands for
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    92
  suitable modes to be asserted.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    93
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
    94
  \<^descr> \<^ML>\<open>Proof.simple_goal\<close>~\<open>state\<close> returns the structured Isar goal (if
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    95
  available) in the form seen by ``simple'' methods (like @{method simp} or
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
    96
  @{method blast}). The Isar goal facts are already inserted as premises into
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
    97
  the subgoals, which are presented individually as in \<^ML>\<open>Proof.goal\<close>.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
    98
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
    99
  \<^descr> \<^ML>\<open>Proof.goal\<close>~\<open>state\<close> returns the structured Isar goal (if available)
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   100
  in the form seen by regular methods (like @{method rule}). The auxiliary
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   101
  internal encoding of Pure conjunctions is split into individual subgoals as
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   102
  usual.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
   103
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   104
  \<^descr> \<^ML>\<open>Proof.raw_goal\<close>~\<open>state\<close> returns the structured Isar goal (if
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   105
  available) in the raw internal form seen by ``raw'' methods (like @{method
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   106
  induct}). This form is rarely appropriate for diagnostic tools; \<^ML>\<open>Proof.simple_goal\<close> or \<^ML>\<open>Proof.goal\<close> should be used in most situations.
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
   107
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   108
  \<^descr> \<^ML>\<open>Proof.theorem\<close>~\<open>before_qed after_qed statement ctxt\<close> initializes a
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   109
  toplevel Isar proof state within a given context.
39845
50f42116ebdb more on Proof.theorem;
wenzelm
parents: 39844
diff changeset
   110
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   111
  The optional \<open>before_qed\<close> method is applied at the end of the proof, just
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   112
  before extracting the result (this feature is rarely used).
39845
50f42116ebdb more on Proof.theorem;
wenzelm
parents: 39844
diff changeset
   113
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   114
  The \<open>after_qed\<close> continuation receives the extracted result in order to apply
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   115
  it to the final context in a suitable way (e.g.\ storing named facts). Note
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   116
  that at this generic level the target context is specified as \<^ML_type>\<open>Proof.context\<close>, but the usual wrapping of toplevel proofs into command
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   117
  transactions will provide a \<^ML_type>\<open>local_theory\<close> here
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   118
  (\chref{ch:local-theory}). This affects the way how results are stored.
39845
50f42116ebdb more on Proof.theorem;
wenzelm
parents: 39844
diff changeset
   119
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   120
  The \<open>statement\<close> is given as a nested list of terms, each associated with
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   121
  optional @{keyword "is"} patterns as usual in the Isar source language. The
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   122
  original nested list structure over terms is turned into one over theorems
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   123
  when \<open>after_qed\<close> is invoked.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   124
\<close>
39842
7205191afde4 more on "Isar language elements";
wenzelm
parents: 30272
diff changeset
   125
20520
wenzelm
parents: 20472
diff changeset
   126
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   127
text %mlantiq \<open>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   128
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   129
  @{ML_antiquotation_def "Isar.goal"} & : & \<open>ML_antiquotation\<close> \\
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   130
  \end{matharray}
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   131
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   132
  \<^descr> \<open>@{Isar.goal}\<close> refers to the regular goal state (if available) of the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   133
  current proof state managed by the Isar toplevel --- as abstract value.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   134
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   135
  This only works for diagnostic ML commands, such as @{command ML_val} or
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   136
  @{command ML_command}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   137
\<close>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   138
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   139
text %mlex \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   140
  The following example peeks at a certain goal configuration.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   141
\<close>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   142
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   143
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   144
begin
39846
cb6634eb8926 examples in Isabelle/HOL;
wenzelm
parents: 39845
diff changeset
   145
  have A and B and C
58728
42398b610f86 tuned spacing;
wenzelm
parents: 58618
diff changeset
   146
    ML_val
42398b610f86 tuned spacing;
wenzelm
parents: 58618
diff changeset
   147
     \<open>val n = Thm.nprems_of (#goal @{Isar.goal});
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   148
      \<^assert> (n = 3);\<close>
58801
f420225a22d6 'notepad' requires proper nesting of begin/end;
wenzelm
parents: 58728
diff changeset
   149
    sorry
f420225a22d6 'notepad' requires proper nesting of begin/end;
wenzelm
parents: 58728
diff changeset
   150
end
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   151
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   152
text \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   153
  Here we see 3 individual subgoals in the same way as regular proof methods
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   154
  would do.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   155
\<close>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   156
20520
wenzelm
parents: 20472
diff changeset
   157
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   158
section \<open>Proof methods\<close>
20472
wenzelm
parents:
diff changeset
   159
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   160
text \<open>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   161
  A \<open>method\<close> is a function \<open>thm\<^sup>* \<rightarrow> context * goal \<rightarrow> (context \<times> goal)\<^sup>*\<^sup>*\<close>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   162
  that operates on the full Isar goal configuration with context, goal facts,
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   163
  and tactical goal state and enumerates possible follow-up goal states. Under
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   164
  normal circumstances, the goal context remains unchanged, but it is also
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   165
  possible to declare named extensions of the proof context (\<^emph>\<open>cases\<close>).
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   166
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   167
  This means a proof method is like a structurally enhanced tactic (cf.\
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   168
  \secref{sec:tactics}). The well-formedness conditions for tactics need to
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   169
  hold for methods accordingly, with the following additions.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   170
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   171
  \<^item> Goal addressing is further limited either to operate uniformly on \<^emph>\<open>all\<close>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   172
  subgoals, or specifically on the \<^emph>\<open>first\<close> subgoal.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   173
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   174
  Exception: old-style tactic emulations that are embedded into the method
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   175
  space, e.g.\ @{method rule_tac}.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   176
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   177
  \<^item> A non-trivial method always needs to make progress: an identical follow-up
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   178
  goal state has to be avoided.\<^footnote>\<open>This enables the user to write method
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   179
  expressions like \<open>meth\<^sup>+\<close> without looping, while the trivial do-nothing case
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   180
  can be recovered via \<open>meth\<^sup>?\<close>.\<close>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   181
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   182
  Exception: trivial stuttering steps, such as ``@{method -}'' or @{method
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   183
  succeed}.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   184
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   185
  \<^item> Goal facts passed to the method must not be ignored. If there is no
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   186
  sensible use of facts outside the goal state, facts should be inserted into
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   187
  the subgoals that are addressed by the method.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   188
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   189
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   190
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   191
  Syntactically, the language of proof methods appears as arguments to Isar
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   192
  commands like @{command "by"} or @{command apply}. User-space additions are
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   193
  reasonably easy by plugging suitable method-valued parser functions into the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   194
  framework, using the @{command method_setup} command, for example.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   195
39844
wenzelm
parents: 39843
diff changeset
   196
  To get a better idea about the range of possibilities, consider the
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   197
  following Isar proof schemes. This is the general form of structured proof
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   198
  text:
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   199
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   200
  \<^medskip>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   201
  \begin{tabular}{l}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   202
  @{command from}~\<open>facts\<^sub>1\<close>~@{command have}~\<open>props\<close>~@{command using}~\<open>facts\<^sub>2\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   203
  @{command proof}~\<open>(initial_method)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   204
  \quad\<open>body\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   205
  @{command qed}~\<open>(terminal_method)\<close> \\
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   206
  \end{tabular}
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   207
  \<^medskip>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   208
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   209
  The goal configuration consists of \<open>facts\<^sub>1\<close> and \<open>facts\<^sub>2\<close> appended in that
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   210
  order, and various \<open>props\<close> being claimed. The \<open>initial_method\<close> is invoked
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   211
  with facts and goals together and refines the problem to something that is
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   212
  handled recursively in the proof \<open>body\<close>. The \<open>terminal_method\<close> has another
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   213
  chance to finish any remaining subgoals, but it does not see the facts of
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   214
  the initial step.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   215
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   216
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   217
  This pattern illustrates unstructured proof scripts:
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   218
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   219
  \<^medskip>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   220
  \begin{tabular}{l}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   221
  @{command have}~\<open>props\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   222
  \quad@{command using}~\<open>facts\<^sub>1\<close>~@{command apply}~\<open>method\<^sub>1\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   223
  \quad@{command apply}~\<open>method\<^sub>2\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   224
  \quad@{command using}~\<open>facts\<^sub>3\<close>~@{command apply}~\<open>method\<^sub>3\<close> \\
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   225
  \quad@{command done} \\
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   226
  \end{tabular}
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   227
  \<^medskip>
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   228
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   229
  The \<open>method\<^sub>1\<close> operates on the original claim while using \<open>facts\<^sub>1\<close>. Since
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   230
  the @{command apply} command structurally resets the facts, the \<open>method\<^sub>2\<close>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   231
  will operate on the remaining goal state without facts. The \<open>method\<^sub>3\<close> will
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   232
  see again a collection of \<open>facts\<^sub>3\<close> that has been inserted into the script
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   233
  explicitly.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   234
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   235
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   236
  Empirically, any Isar proof method can be categorized as follows.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   237
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   238
    \<^enum> \<^emph>\<open>Special method with cases\<close> with named context additions associated
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   239
    with the follow-up goal state.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   240
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   241
    Example: @{method "induct"}, which is also a ``raw'' method since it
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   242
    operates on the internal representation of simultaneous claims as Pure
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   243
    conjunction (\secref{sec:logic-aux}), instead of separate subgoals
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   244
    (\secref{sec:tactical-goals}).
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   245
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   246
    \<^enum> \<^emph>\<open>Structured method\<close> with strong emphasis on facts outside the goal
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   247
    state.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   248
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   249
    Example: @{method "rule"}, which captures the key ideas behind structured
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   250
    reasoning in Isar in its purest form.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   251
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   252
    \<^enum> \<^emph>\<open>Simple method\<close> with weaker emphasis on facts, which are inserted into
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   253
    subgoals to emulate old-style tactical ``premises''.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   254
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   255
    Examples: @{method "simp"}, @{method "blast"}, @{method "auto"}.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   256
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   257
    \<^enum> \<^emph>\<open>Old-style tactic emulation\<close> with detailed numeric goal addressing and
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   258
    explicit references to entities of the internal goal state (which are
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   259
    otherwise invisible from proper Isar proof text). The naming convention
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   260
    \<open>foo_tac\<close> makes this special non-standard status clear.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   261
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   262
    Example: @{method "rule_tac"}.
39843
21d189bfdfd1 more examples;
wenzelm
parents: 39842
diff changeset
   263
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   264
  When implementing proof methods, it is advisable to study existing
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   265
  implementations carefully and imitate the typical ``boiler plate'' for
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   266
  context-sensitive parsing and further combinators to wrap-up tactic
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   267
  expressions as methods.\<^footnote>\<open>Aliases or abbreviations of the standard method
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   268
  combinators should be avoided. Note that from Isabelle99 until Isabelle2009
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   269
  the system did provide various odd combinations of method syntax wrappers
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   270
  that made applications more complicated than necessary.\<close>
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   271
\<close>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   272
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   273
text %mlref \<open>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   274
  \begin{mldecls}
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   275
  @{index_ML_type Proof.method} \\
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   276
  @{index_ML CONTEXT_METHOD: "(thm list -> context_tactic) -> Proof.method"} \\
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   277
  @{index_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   278
  @{index_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   279
  @{index_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   280
  @{index_ML Method.insert_tac: "Proof.context -> thm list -> int -> tactic"} \\
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   281
  @{index_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser ->
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   282
  string -> theory -> theory"} \\
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   283
  \end{mldecls}
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   284
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   285
  \<^descr> Type \<^ML_type>\<open>Proof.method\<close> represents proof methods as abstract type.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   286
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   287
  \<^descr> \<^ML>\<open>CONTEXT_METHOD\<close>~\<open>(fn facts => context_tactic)\<close> wraps \<open>context_tactic\<close>
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   288
  depending on goal facts as a general proof method that may change the proof
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   289
  context dynamically. A typical operation is \<^ML>\<open>Proof_Context.update_cases\<close>, which is wrapped up as combinator @{index_ML
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61656
diff changeset
   290
  CONTEXT_CASES} for convenience.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   291
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   292
  \<^descr> \<^ML>\<open>METHOD\<close>~\<open>(fn facts => tactic)\<close> wraps \<open>tactic\<close> depending on goal facts
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   293
  as regular proof method; the goal context is passed via method syntax.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   294
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   295
  \<^descr> \<^ML>\<open>SIMPLE_METHOD\<close>~\<open>tactic\<close> wraps a tactic that addresses all subgoals
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   296
  uniformly as simple proof method. Goal facts are already inserted into all
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   297
  subgoals before \<open>tactic\<close> is applied.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   298
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   299
  \<^descr> \<^ML>\<open>SIMPLE_METHOD'\<close>~\<open>tactic\<close> wraps a tactic that addresses a specific
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   300
  subgoal as simple proof method that operates on subgoal 1. Goal facts are
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   301
  inserted into the subgoal then the \<open>tactic\<close> is applied.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   302
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   303
  \<^descr> \<^ML>\<open>Method.insert_tac\<close>~\<open>ctxt facts i\<close> inserts \<open>facts\<close> into subgoal \<open>i\<close>.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   304
  This is convenient to reproduce part of the \<^ML>\<open>SIMPLE_METHOD\<close> or \<^ML>\<open>SIMPLE_METHOD'\<close> wrapping within regular \<^ML>\<open>METHOD\<close>, for example.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   305
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   306
  \<^descr> \<^ML>\<open>Method.setup\<close>~\<open>name parser description\<close> provides the functionality of
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   307
  the Isar command @{command method_setup} as ML function.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   308
\<close>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   309
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   310
text %mlex \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   311
  See also @{command method_setup} in @{cite "isabelle-isar-ref"} which
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   312
  includes some abstract examples.
39851
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   313
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   314
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   315
  The following toy examples illustrate how the goal facts and state are
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   316
  passed to proof methods. The predefined proof method called ``@{method
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   317
  tactic}'' wraps ML source of type \<^ML_type>\<open>tactic\<close> (abstracted over
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   318
  \<^ML_text>\<open>facts\<close>). This allows immediate experimentation without parsing of
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   319
  concrete syntax.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   320
\<close>
20472
wenzelm
parents:
diff changeset
   321
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   322
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   323
begin
57340
wenzelm
parents: 56579
diff changeset
   324
  fix A B :: bool
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   325
  assume a: A and b: B
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   326
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   327
  have "A \<and> B"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   328
    apply (tactic \<open>resolve_tac \<^context> @{thms conjI} 1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   329
    using a apply (tactic \<open>resolve_tac \<^context> facts 1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   330
    using b apply (tactic \<open>resolve_tac \<^context> facts 1\<close>)
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   331
    done
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   332
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   333
  have "A \<and> B"
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   334
    using a and b
57340
wenzelm
parents: 56579
diff changeset
   335
    ML_val \<open>@{Isar.goal}\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   336
    apply (tactic \<open>Method.insert_tac \<^context> facts 1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   337
    apply (tactic \<open>(resolve_tac \<^context> @{thms conjI} THEN_ALL_NEW assume_tac \<^context>) 1\<close>)
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   338
    done
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   339
end
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   340
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   341
text \<open>
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   342
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   343
  The next example implements a method that simplifies the first subgoal by
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   344
  rewrite rules that are given as arguments.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   345
\<close>
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   346
57340
wenzelm
parents: 56579
diff changeset
   347
method_setup my_simp =
wenzelm
parents: 56579
diff changeset
   348
  \<open>Attrib.thms >> (fn thms => fn ctxt =>
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   349
    SIMPLE_METHOD' (fn i =>
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   350
      CHANGED (asm_full_simp_tac
57340
wenzelm
parents: 56579
diff changeset
   351
        (put_simpset HOL_basic_ss ctxt addsimps thms) i)))\<close>
wenzelm
parents: 56579
diff changeset
   352
  "rewrite subgoal by given rules"
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   353
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   354
text \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   355
  The concrete syntax wrapping of @{command method_setup} always
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   356
  passes-through the proof context at the end of parsing, but it is not used
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   357
  in this example.
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   358
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   359
  The \<^ML>\<open>Attrib.thms\<close> parser produces a list of theorems from the usual Isar
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   360
  syntax involving attribute expressions etc.\ (syntax category @{syntax
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   361
  thms}) @{cite "isabelle-isar-ref"}. The resulting \<^ML_text>\<open>thms\<close> are
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   362
  added to \<^ML>\<open>HOL_basic_ss\<close> which already contains the basic Simplifier
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   363
  setup for HOL.
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   364
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   365
  The tactic \<^ML>\<open>asm_full_simp_tac\<close> is the one that is also used in method
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   366
  @{method simp} by default. The extra wrapping by the \<^ML>\<open>CHANGED\<close> tactical
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   367
  ensures progress of simplification: identical goal states are filtered out
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   368
  explicitly to make the raw tactic conform to standard Isar method behaviour.
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   369
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   370
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   371
  Method @{method my_simp} can be used in Isar proofs like this:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   372
\<close>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   373
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   374
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   375
begin
57340
wenzelm
parents: 56579
diff changeset
   376
  fix a b c :: 'a
39851
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   377
  assume a: "a = b"
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   378
  assume b: "b = c"
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   379
  have "a = c" by (my_simp a b)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   380
end
39851
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   381
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   382
text \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   383
  Here is a similar method that operates on all subgoals, instead of just the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   384
  first one.\<close>
39851
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   385
57340
wenzelm
parents: 56579
diff changeset
   386
method_setup my_simp_all =
wenzelm
parents: 56579
diff changeset
   387
  \<open>Attrib.thms >> (fn thms => fn ctxt =>
39851
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   388
    SIMPLE_METHOD
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   389
      (CHANGED
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   390
        (ALLGOALS (asm_full_simp_tac
57340
wenzelm
parents: 56579
diff changeset
   391
          (put_simpset HOL_basic_ss ctxt addsimps thms)))))\<close>
wenzelm
parents: 56579
diff changeset
   392
  "rewrite all subgoals by given rules"
39851
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   393
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   394
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   395
begin
57340
wenzelm
parents: 56579
diff changeset
   396
  fix a b c :: 'a
39851
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   397
  assume a: "a = b"
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   398
  assume b: "b = c"
7219a771ab63 more examples;
wenzelm
parents: 39849
diff changeset
   399
  have "a = c" and "c = b" by (my_simp_all a b)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   400
end
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   401
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   402
text \<open>
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   403
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   404
  Apart from explicit arguments, common proof methods typically work with a
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   405
  default configuration provided by the context. As a shortcut to rule
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   406
  management we use a cheap solution via the @{command named_theorems} command
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   407
  to declare a dynamic fact in the context.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   408
\<close>
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   409
57946
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   410
named_theorems my_simp
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   411
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   412
text \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   413
  The proof method is now defined as before, but we append the explicit
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   414
  arguments and the rules from the context.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   415
\<close>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   416
57340
wenzelm
parents: 56579
diff changeset
   417
method_setup my_simp' =
wenzelm
parents: 56579
diff changeset
   418
  \<open>Attrib.thms >> (fn thms => fn ctxt =>
57946
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   419
    let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   420
      val my_simps = Named_Theorems.get ctxt \<^named_theorems>\<open>my_simp\<close>
57946
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   421
    in
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   422
      SIMPLE_METHOD' (fn i =>
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   423
        CHANGED (asm_full_simp_tac
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   424
          (put_simpset HOL_basic_ss ctxt
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   425
            addsimps (thms @ my_simps)) i))
6a26aa5fa65e updated documentation concerning 'named_theorems';
wenzelm
parents: 57342
diff changeset
   426
    end)\<close>
57340
wenzelm
parents: 56579
diff changeset
   427
  "rewrite subgoal by given rules and my_simp rules from the context"
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   428
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   429
text \<open>
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   430
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   431
  Method @{method my_simp'} can be used in Isar proofs like this:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   432
\<close>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   433
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   434
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   435
begin
57342
wenzelm
parents: 57340
diff changeset
   436
  fix a b c :: 'a
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   437
  assume [my_simp]: "a \<equiv> b"
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   438
  assume [my_simp]: "b \<equiv> c"
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   439
  have "a \<equiv> c" by my_simp'
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40800
diff changeset
   440
end
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   441
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   442
text \<open>
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   443
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   444
  The @{method my_simp} variants defined above are ``simple'' methods, i.e.\
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   445
  the goal facts are merely inserted as goal premises by the \<^ML>\<open>SIMPLE_METHOD'\<close> or \<^ML>\<open>SIMPLE_METHOD\<close> wrapper. For proof methods that are
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   446
  similar to the standard collection of @{method simp}, @{method blast},
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   447
  @{method fast}, @{method auto} there is little more that can be done.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   448
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   449
  Note that using the primary goal facts in the same manner as the method
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   450
  arguments obtained via concrete syntax or the context does not meet the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   451
  requirement of ``strong emphasis on facts'' of regular proof methods,
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   452
  because rewrite rules as used above can be easily ignored. A proof text
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   453
  ``@{command using}~\<open>foo\<close>~@{command "by"}~\<open>my_simp\<close>'' where \<open>foo\<close> is not used
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   454
  would deceive the reader.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   455
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 60754
diff changeset
   456
  \<^medskip>
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   457
  The technical treatment of rules from the context requires further
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   458
  attention. Above we rebuild a fresh \<^ML_type>\<open>simpset\<close> from the arguments
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   459
  and \<^emph>\<open>all\<close> rules retrieved from the context on every invocation of the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   460
  method. This does not scale to really large collections of rules, which
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   461
  easily emerges in the context of a big theory library, for example.
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   462
39848
fc88b943e1b2 misc tuning and clarification;
wenzelm
parents: 39847
diff changeset
   463
  This is an inherent limitation of the simplistic rule management via
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   464
  @{command named_theorems}, because it lacks tool-specific storage and
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   465
  retrieval. More realistic applications require efficient index-structures
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   466
  that organize theorems in a customized manner, such as a discrimination net
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   467
  that is indexed by the left-hand sides of rewrite rules. For variations on
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   468
  the Simplifier, re-use of the existing type \<^ML_type>\<open>simpset\<close> is adequate,
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   469
  but scalability would require it be maintained statically within the context
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   470
  data, not dynamically on each tool invocation.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   471
\<close>
39847
da8c3fc5e314 more on "Proof methods";
wenzelm
parents: 39846
diff changeset
   472
29759
bcb79ddf57da removed rudiments of glossary;
wenzelm
parents: 29755
diff changeset
   473
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   474
section \<open>Attributes \label{sec:attributes}\<close>
20472
wenzelm
parents:
diff changeset
   475
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   476
text \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   477
  An \<^emph>\<open>attribute\<close> is a function \<open>context \<times> thm \<rightarrow> context \<times> thm\<close>, which means
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   478
  both a (generic) context and a theorem can be modified simultaneously. In
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   479
  practice this mixed form is very rare, instead attributes are presented
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   480
  either as \<^emph>\<open>declaration attribute:\<close> \<open>thm \<rightarrow> context \<rightarrow> context\<close> or \<^emph>\<open>rule
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   481
  attribute:\<close> \<open>context \<rightarrow> thm \<rightarrow> thm\<close>.
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   482
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   483
  Attributes can have additional arguments via concrete syntax. There is a
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   484
  collection of context-sensitive parsers for various logical entities (types,
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   485
  terms, theorems). These already take care of applying morphisms to the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   486
  arguments when attribute expressions are moved into a different context (see
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   487
  also \secref{sec:morphisms}).
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   488
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   489
  When implementing declaration attributes, it is important to operate exactly
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   490
  on the variant of the generic context that is provided by the system, which
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   491
  is either global theory context or local proof context. In particular, the
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   492
  background theory of a local context must not be modified in this
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   493
  situation!
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   494
\<close>
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   495
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   496
text %mlref \<open>
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   497
  \begin{mldecls}
45414
8ca612982014 updated functor Named_Thms;
wenzelm
parents: 40964
diff changeset
   498
  @{index_ML_type attribute} \\
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   499
  @{index_ML Thm.rule_attribute: "thm list ->
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   500
  (Context.generic -> thm -> thm) -> attribute"} \\
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   501
  @{index_ML Thm.declaration_attribute: "
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   502
  (thm -> Context.generic -> Context.generic) -> attribute"} \\
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   503
  @{index_ML Attrib.setup: "binding -> attribute context_parser ->
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   504
  string -> theory -> theory"} \\
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   505
  \end{mldecls}
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   506
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   507
  \<^descr> Type \<^ML_type>\<open>attribute\<close> represents attributes as concrete type alias.
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   508
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   509
  \<^descr> \<^ML>\<open>Thm.rule_attribute\<close>~\<open>thms (fn context => rule)\<close> wraps a
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   510
  context-dependent rule (mapping on \<^ML_type>\<open>thm\<close>) as attribute.
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   511
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   512
  The \<open>thms\<close> are additional parameters: when forming an abstract closure, the
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   513
  system may provide dummy facts that are propagated according to strict
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   514
  evaluation discipline. In that case, \<open>rule\<close> is bypassed.
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   515
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   516
  \<^descr> \<^ML>\<open>Thm.declaration_attribute\<close>~\<open>(fn thm => decl)\<close> wraps a
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   517
  theorem-dependent declaration (mapping on \<^ML_type>\<open>Context.generic\<close>) as
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   518
  attribute.
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   519
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   520
  When forming an abstract closure, the system may provide a dummy fact as
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61841
diff changeset
   521
  \<open>thm\<close>. In that case, \<open>decl\<close> is bypassed.
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   522
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   523
  \<^descr> \<^ML>\<open>Attrib.setup\<close>~\<open>name parser description\<close> provides the functionality of
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   524
  the Isar command @{command attribute_setup} as ML function.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   525
\<close>
39849
b7b1a9e8f7c2 more on "Attributes";
wenzelm
parents: 39848
diff changeset
   526
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   527
text %mlantiq \<open>
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45414
diff changeset
   528
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   529
  @{ML_antiquotation_def attributes} & : & \<open>ML_antiquotation\<close> \\
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45414
diff changeset
   530
  \end{matharray}
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45414
diff changeset
   531
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   532
  \<^rail>\<open>
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45414
diff changeset
   533
  @@{ML_antiquotation attributes} attributes
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   534
  \<close>
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45414
diff changeset
   535
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   536
  \<^descr> \<open>@{attributes [\<dots>]}\<close> embeds attribute source representation into the ML
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   537
  text, which is particularly useful with declarations like \<^ML>\<open>Local_Theory.note\<close>. Attribute names are internalized at compile time, but
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   538
  the source is unevaluated. This means attributes with formal arguments
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   539
  (types, terms, theorems) may be subject to odd effects of dynamic scoping!
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   540
\<close>
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45414
diff changeset
   541
61854
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   542
text %mlex \<open>
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   543
  See also @{command attribute_setup} in @{cite "isabelle-isar-ref"} which
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   544
  includes some abstract examples.
38b049cd3aad tuned whitespace;
wenzelm
parents: 61853
diff changeset
   545
\<close>
30272
2d612824e642 regenerated document;
wenzelm
parents: 30270
diff changeset
   546
20472
wenzelm
parents:
diff changeset
   547
end