doc-src/IsarRef/Thy/Generic.thy
author wenzelm
Sun, 15 May 2011 17:06:35 +0200
changeset 42813 6c841fa92fa2
parent 42705 528a2ba8fa74
child 42925 c6c4f997ad87
permissions -rw-r--r--
optional description for 'attribute_setup' and 'method_setup';
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     1
theory Generic
42651
e3fdb7c96be5 formal Base theory;
wenzelm
parents: 42626
diff changeset
     2
imports Base Main
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     3
begin
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     4
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     5
chapter {* Generic tools and packages \label{ch:gen-tools} *}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     6
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
     7
section {* Configuration options \label{sec:config} *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     8
40291
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
     9
text {* Isabelle/Pure maintains a record of named configuration
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    10
  options within the theory or proof context, with values of type
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    11
  @{ML_type bool}, @{ML_type int}, @{ML_type real}, or @{ML_type
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    12
  string}.  Tools may declare options in ML, and then refer to these
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    13
  values (relative to the context).  Thus global reference variables
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    14
  are easily avoided.  The user may change the value of a
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    15
  configuration option by means of an associated attribute of the same
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    16
  name.  This form of context declaration works particularly well with
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    17
  commands such as @{command "declare"} or @{command "using"} like
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    18
  this:
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    19
*}
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    20
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    21
declare [[show_main_goal = false]]
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    22
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    23
notepad
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    24
begin
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    25
  note [[show_main_goal = true]]
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    26
end
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    27
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    28
text {* For historical reasons, some tools cannot take the full proof
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    29
  context into account and merely refer to the background theory.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    30
  This is accommodated by configuration options being declared as
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    31
  ``global'', which may not be changed within a local context.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    32
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    33
  \begin{matharray}{rcll}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    34
    @{command_def "print_configs"} & : & @{text "context \<rightarrow>"} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    35
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    36
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    37
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    38
    @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    39
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    40
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    41
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    42
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    43
  \item @{command "print_configs"} prints the available configuration
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    44
  options, with names, types, and current values.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    45
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    46
  \item @{text "name = value"} as an attribute expression modifies the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    47
  named option, with the syntax of the value depending on the option's
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    48
  type.  For @{ML_type bool} the default value is @{text true}.  Any
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    49
  attempt to change a global option in a local context is ignored.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    50
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    51
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    52
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    53
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    54
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
    55
section {* Basic proof tools *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    56
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    57
subsection {* Miscellaneous methods and attributes \label{sec:misc-meth-att} *}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    58
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    59
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    60
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    61
    @{method_def unfold} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    62
    @{method_def fold} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    63
    @{method_def insert} & : & @{text method} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    64
    @{method_def erule}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    65
    @{method_def drule}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    66
    @{method_def frule}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    67
    @{method_def succeed} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    68
    @{method_def fail} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    69
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    70
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    71
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    72
    (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thmrefs}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    73
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    74
    (@@{method erule} | @@{method drule} | @@{method frule})
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    75
      ('(' @{syntax nat} ')')? @{syntax thmrefs}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    76
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    77
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    78
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    79
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    80
  \item @{method unfold}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} and @{method fold}~@{text
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    81
  "a\<^sub>1 \<dots> a\<^sub>n"} expand (or fold back) the given definitions throughout
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    82
  all goals; any chained facts provided are inserted into the goal and
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    83
  subject to rewriting as well.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    84
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    85
  \item @{method insert}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} inserts theorems as facts
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    86
  into all goals of the proof state.  Note that current facts
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    87
  indicated for forward chaining are ignored.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    88
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    89
  \item @{method erule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, @{method
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    90
  drule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, and @{method frule}~@{text
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    91
  "a\<^sub>1 \<dots> a\<^sub>n"} are similar to the basic @{method rule}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    92
  method (see \secref{sec:pure-meth-att}), but apply rules by
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    93
  elim-resolution, destruct-resolution, and forward-resolution,
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    94
  respectively \cite{isabelle-implementation}.  The optional natural
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    95
  number argument (default 0) specifies additional assumption steps to
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    96
  be performed here.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    97
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    98
  Note that these methods are improper ones, mainly serving for
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    99
  experimentation and tactic script emulation.  Different modes of
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   100
  basic rule application are usually expressed in Isar at the proof
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   101
  language level, rather than via implicit proof state manipulations.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   102
  For example, a proper single-step elimination would be done using
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   103
  the plain @{method rule} method, with forward chaining of current
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   104
  facts.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   105
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   106
  \item @{method succeed} yields a single (unchanged) result; it is
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   107
  the identity of the ``@{text ","}'' method combinator (cf.\
28754
6f2e67a3dfaa moved section "Proof method expressions" to proof chapter;
wenzelm
parents: 27248
diff changeset
   108
  \secref{sec:proof-meth}).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   109
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   110
  \item @{method fail} yields an empty result sequence; it is the
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   111
  identity of the ``@{text "|"}'' method combinator (cf.\
28754
6f2e67a3dfaa moved section "Proof method expressions" to proof chapter;
wenzelm
parents: 27248
diff changeset
   112
  \secref{sec:proof-meth}).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   113
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   114
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   115
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   116
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   117
    @{attribute_def tagged} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   118
    @{attribute_def untagged} & : & @{text attribute} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   119
    @{attribute_def THEN} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   120
    @{attribute_def COMP} & : & @{text attribute} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   121
    @{attribute_def unfolded} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   122
    @{attribute_def folded} & : & @{text attribute} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   123
    @{attribute_def rotated} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   124
    @{attribute_def (Pure) elim_format} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   125
    @{attribute_def standard}@{text "\<^sup>*"} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   126
    @{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   127
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   128
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   129
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   130
    @@{attribute tagged} @{syntax name} @{syntax name}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   131
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   132
    @@{attribute untagged} @{syntax name}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   133
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   134
    (@@{attribute THEN} | @@{attribute COMP}) ('[' @{syntax nat} ']')? @{syntax thmref}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   135
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   136
    (@@{attribute unfolded} | @@{attribute folded}) @{syntax thmrefs}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   137
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   138
    @@{attribute rotated} @{syntax int}?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   139
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   140
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   141
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   142
28764
b65194fe4434 fixed/tuned syntax for attribute "tagged";
wenzelm
parents: 28761
diff changeset
   143
  \item @{attribute tagged}~@{text "name value"} and @{attribute
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   144
  untagged}~@{text name} add and remove \emph{tags} of some theorem.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   145
  Tags may be any list of string pairs that serve as formal comment.
28764
b65194fe4434 fixed/tuned syntax for attribute "tagged";
wenzelm
parents: 28761
diff changeset
   146
  The first string is considered the tag name, the second its value.
b65194fe4434 fixed/tuned syntax for attribute "tagged";
wenzelm
parents: 28761
diff changeset
   147
  Note that @{attribute untagged} removes any tags of the same name.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   148
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   149
  \item @{attribute THEN}~@{text a} and @{attribute COMP}~@{text a}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   150
  compose rules by resolution.  @{attribute THEN} resolves with the
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   151
  first premise of @{text a} (an alternative position may be also
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   152
  specified); the @{attribute COMP} version skips the automatic
30462
wenzelm
parents: 30397
diff changeset
   153
  lifting process that is normally intended (cf.\ @{ML "op RS"} and
wenzelm
parents: 30397
diff changeset
   154
  @{ML "op COMP"} in \cite{isabelle-implementation}).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   155
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   156
  \item @{attribute unfolded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} and @{attribute
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   157
  folded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} expand and fold back again the given
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   158
  definitions throughout a rule.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   159
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   160
  \item @{attribute rotated}~@{text n} rotate the premises of a
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   161
  theorem by @{text n} (default 1).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   162
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   163
  \item @{attribute (Pure) elim_format} turns a destruction rule into
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   164
  elimination rule format, by resolving with the rule @{prop "PROP A \<Longrightarrow>
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   165
  (PROP A \<Longrightarrow> PROP B) \<Longrightarrow> PROP B"}.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   166
  
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   167
  Note that the Classical Reasoner (\secref{sec:classical}) provides
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   168
  its own version of this operation.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   169
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   170
  \item @{attribute standard} puts a theorem into the standard form of
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   171
  object-rules at the outermost theory level.  Note that this
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   172
  operation violates the local proof context (including active
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   173
  locales).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   174
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   175
  \item @{attribute no_vars} replaces schematic variables by free
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   176
  ones; this is mainly for tuning output of pretty printed theorems.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   177
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   178
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   179
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   180
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   181
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   182
subsection {* Low-level equational reasoning *}
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   183
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   184
text {*
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   185
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   186
    @{method_def subst} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   187
    @{method_def hypsubst} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   188
    @{method_def split} & : & @{text method} \\
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   189
  \end{matharray}
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   190
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   191
  @{rail "
42704
3f19e324ff59 tuned rail diagrams and layout;
wenzelm
parents: 42655
diff changeset
   192
    @@{method subst} ('(' 'asm' ')')? \\ ('(' (@{syntax nat}+) ')')? @{syntax thmref}
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   193
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   194
    @@{method split} ('(' 'asm' ')')? @{syntax thmrefs}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   195
  "}
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   196
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   197
  These methods provide low-level facilities for equational reasoning
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   198
  that are intended for specialized applications only.  Normally,
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   199
  single step calculations would be performed in a structured text
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   200
  (see also \secref{sec:calculation}), while the Simplifier methods
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   201
  provide the canonical way for automated normalization (see
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   202
  \secref{sec:simplifier}).
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   203
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   204
  \begin{description}
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   205
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   206
  \item @{method subst}~@{text eq} performs a single substitution step
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   207
  using rule @{text eq}, which may be either a meta or object
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   208
  equality.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   209
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   210
  \item @{method subst}~@{text "(asm) eq"} substitutes in an
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   211
  assumption.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   212
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   213
  \item @{method subst}~@{text "(i \<dots> j) eq"} performs several
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   214
  substitutions in the conclusion. The numbers @{text i} to @{text j}
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   215
  indicate the positions to substitute at.  Positions are ordered from
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   216
  the top of the term tree moving down from left to right. For
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   217
  example, in @{text "(a + b) + (c + d)"} there are three positions
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   218
  where commutativity of @{text "+"} is applicable: 1 refers to @{text
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   219
  "a + b"}, 2 to the whole term, and 3 to @{text "c + d"}.
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   220
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   221
  If the positions in the list @{text "(i \<dots> j)"} are non-overlapping
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   222
  (e.g.\ @{text "(2 3)"} in @{text "(a + b) + (c + d)"}) you may
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   223
  assume all substitutions are performed simultaneously.  Otherwise
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   224
  the behaviour of @{text subst} is not specified.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   225
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   226
  \item @{method subst}~@{text "(asm) (i \<dots> j) eq"} performs the
27071
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   227
  substitutions in the assumptions. The positions refer to the
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   228
  assumptions in order from left to right.  For example, given in a
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   229
  goal of the form @{text "P (a + b) \<Longrightarrow> P (c + d) \<Longrightarrow> \<dots>"}, position 1 of
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   230
  commutativity of @{text "+"} is the subterm @{text "a + b"} and
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   231
  position 2 is the subterm @{text "c + d"}.
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   232
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   233
  \item @{method hypsubst} performs substitution using some
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   234
  assumption; this only works for equations of the form @{text "x =
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   235
  t"} where @{text x} is a free or bound variable.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   236
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   237
  \item @{method split}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} performs single-step case
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   238
  splitting using the given rules.  By default, splitting is performed
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   239
  in the conclusion of a goal; the @{text "(asm)"} option indicates to
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   240
  operate on assumptions instead.
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   241
  
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   242
  Note that the @{method simp} method already involves repeated
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   243
  application of split rules as declared in the current context.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   244
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   245
  \end{description}
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   246
*}
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   247
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   248
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   249
subsection {* Further tactic emulations \label{sec:tactics} *}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   250
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   251
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   252
  The following improper proof methods emulate traditional tactics.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   253
  These admit direct access to the goal state, which is normally
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   254
  considered harmful!  In particular, this may involve both numbered
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   255
  goal addressing (default 1), and dynamic instantiation within the
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   256
  scope of some subgoal.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   257
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   258
  \begin{warn}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   259
    Dynamic instantiations refer to universally quantified parameters
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   260
    of a subgoal (the dynamic context) rather than fixed variables and
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   261
    term abbreviations of a (static) Isar context.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   262
  \end{warn}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   263
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   264
  Tactic emulation methods, unlike their ML counterparts, admit
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   265
  simultaneous instantiation from both dynamic and static contexts.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   266
  If names occur in both contexts goal parameters hide locally fixed
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   267
  variables.  Likewise, schematic variables refer to term
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   268
  abbreviations, if present in the static context.  Otherwise the
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   269
  schematic variable is interpreted as a schematic variable and left
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   270
  to be solved by unification with certain parts of the subgoal.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   271
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   272
  Note that the tactic emulation proof methods in Isabelle/Isar are
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   273
  consistently named @{text foo_tac}.  Note also that variable names
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   274
  occurring on left hand sides of instantiations must be preceded by a
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   275
  question mark if they coincide with a keyword or contain dots.  This
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   276
  is consistent with the attribute @{attribute "where"} (see
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   277
  \secref{sec:pure-meth-att}).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   278
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   279
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   280
    @{method_def rule_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   281
    @{method_def erule_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   282
    @{method_def drule_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   283
    @{method_def frule_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   284
    @{method_def cut_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   285
    @{method_def thin_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   286
    @{method_def subgoal_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   287
    @{method_def rename_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   288
    @{method_def rotate_tac}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   289
    @{method_def tactic}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   290
    @{method_def raw_tactic}@{text "\<^sup>*"} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   291
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   292
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   293
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   294
    (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} |
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
   295
      @@{method frule_tac} | @@{method cut_tac} | @@{method thin_tac}) @{syntax goal_spec}? \\
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   296
    ( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   297
    ;
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
   298
    @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +)
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   299
    ;
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
   300
    @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +)
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   301
    ;
42705
528a2ba8fa74 tuned some syntax names;
wenzelm
parents: 42704
diff changeset
   302
    @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   303
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   304
    (@@{method tactic} | @@{method raw_tactic}) @{syntax text}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   305
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   306
42617
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   307
    dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and')
77d239840285 more precise rail diagrams;
wenzelm
parents: 42596
diff changeset
   308
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   309
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   310
\begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   311
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   312
  \item @{method rule_tac} etc. do resolution of rules with explicit
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   313
  instantiation.  This works the same way as the ML tactics @{ML
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   314
  res_inst_tac} etc. (see \cite{isabelle-implementation})
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   315
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   316
  Multiple rules may be only given if there is no instantiation; then
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   317
  @{method rule_tac} is the same as @{ML resolve_tac} in ML (see
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   318
  \cite{isabelle-implementation}).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   319
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   320
  \item @{method cut_tac} inserts facts into the proof state as
27209
388fd72e9f26 qualified old res_inst_tac variants;
wenzelm
parents: 27092
diff changeset
   321
  assumption of a subgoal, see also @{ML Tactic.cut_facts_tac} in
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   322
  \cite{isabelle-implementation}.  Note that the scope of schematic
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   323
  variables is spread over the main goal statement.  Instantiations
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   324
  may be given as well, see also ML tactic @{ML cut_inst_tac} in
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   325
  \cite{isabelle-implementation}.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   326
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   327
  \item @{method thin_tac}~@{text \<phi>} deletes the specified assumption
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   328
  from a subgoal; note that @{text \<phi>} may contain schematic variables.
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   329
  See also @{ML thin_tac} in \cite{isabelle-implementation}.
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   330
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   331
  \item @{method subgoal_tac}~@{text \<phi>} adds @{text \<phi>} as an
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27223
diff changeset
   332
  assumption to a subgoal.  See also @{ML subgoal_tac} and @{ML
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   333
  subgoals_tac} in \cite{isabelle-implementation}.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   334
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   335
  \item @{method rename_tac}~@{text "x\<^sub>1 \<dots> x\<^sub>n"} renames parameters of a
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   336
  goal according to the list @{text "x\<^sub>1, \<dots>, x\<^sub>n"}, which refers to the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   337
  \emph{suffix} of variables.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   338
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   339
  \item @{method rotate_tac}~@{text n} rotates the assumptions of a
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   340
  goal by @{text n} positions: from right to left if @{text n} is
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   341
  positive, and from left to right if @{text n} is negative; the
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   342
  default value is 1.  See also @{ML rotate_tac} in
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   343
  \cite{isabelle-implementation}.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   344
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   345
  \item @{method tactic}~@{text "text"} produces a proof method from
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   346
  any ML text of type @{ML_type tactic}.  Apart from the usual ML
27223
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   347
  environment and the current proof context, the ML code may refer to
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   348
  the locally bound values @{ML_text facts}, which indicates any
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   349
  current facts used for forward-chaining.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   350
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   351
  \item @{method raw_tactic} is similar to @{method tactic}, but
27223
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   352
  presents the goal state in its raw internal form, where simultaneous
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   353
  subgoals appear as conjunction of the logical framework instead of
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   354
  the usual split into several subgoals.  While feature this is useful
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   355
  for debugging of complex method definitions, it should not never
8546e2407b31 method "tactic": only "facts" as bound value;
wenzelm
parents: 27209
diff changeset
   356
  appear in production theories.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   357
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   358
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   359
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   360
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   361
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   362
section {* The Simplifier \label{sec:simplifier} *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   363
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   364
subsection {* Simplification methods *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   365
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   366
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   367
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   368
    @{method_def simp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   369
    @{method_def simp_all} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   370
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   371
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   372
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   373
    (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   374
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   375
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 35613
diff changeset
   376
    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')'
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   377
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   378
    @{syntax_def simpmod}: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   379
      'split' (() | 'add' | 'del')) ':' @{syntax thmrefs}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   380
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   381
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   382
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   383
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   384
  \item @{method simp} invokes the Simplifier, after declaring
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   385
  additional rules according to the arguments given.  Note that the
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   386
  @{text only} modifier first removes all other rewrite rules,
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   387
  congruences, and looper tactics (including splits), and then behaves
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   388
  like @{text add}.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   389
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   390
  \medskip The @{text cong} modifiers add or delete Simplifier
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   391
  congruence rules (see also \cite{isabelle-ref}), the default is to
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   392
  add.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   393
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   394
  \medskip The @{text split} modifiers add or delete rules for the
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   395
  Splitter (see also \cite{isabelle-ref}), the default is to add.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   396
  This works only if the Simplifier method has been properly setup to
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   397
  include the Splitter (all major object logics such HOL, HOLCF, FOL,
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   398
  ZF do this already).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   399
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   400
  \item @{method simp_all} is similar to @{method simp}, but acts on
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   401
  all goals (backwards from the last to the first one).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   402
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   403
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   404
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   405
  By default the Simplifier methods take local assumptions fully into
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   406
  account, using equational assumptions in the subsequent
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   407
  normalization process, or simplifying assumptions themselves (cf.\
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   408
  @{ML asm_full_simp_tac} in \cite{isabelle-ref}).  In structured
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   409
  proofs this is usually quite well behaved in practice: just the
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   410
  local premises of the actual goal are involved, additional facts may
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   411
  be inserted via explicit forward-chaining (via @{command "then"},
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 30462
diff changeset
   412
  @{command "from"}, @{command "using"} etc.).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   413
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   414
  Additional Simplifier options may be specified to tune the behavior
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   415
  further (mostly for unstructured scripts with many accidental local
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   416
  facts): ``@{text "(no_asm)"}'' means assumptions are ignored
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   417
  completely (cf.\ @{ML simp_tac}), ``@{text "(no_asm_simp)"}'' means
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   418
  assumptions are used in the simplification of the conclusion but are
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   419
  not themselves simplified (cf.\ @{ML asm_simp_tac}), and ``@{text
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   420
  "(no_asm_use)"}'' means assumptions are simplified but are not used
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   421
  in the simplification of each other or the conclusion (cf.\ @{ML
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   422
  full_simp_tac}).  For compatibility reasons, there is also an option
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   423
  ``@{text "(asm_lr)"}'', which means that an assumption is only used
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   424
  for simplifying assumptions which are to the right of it (cf.\ @{ML
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   425
  asm_lr_simp_tac}).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   426
27092
3d79bbdaf2ef simp: depth_limit is now a configuration option;
wenzelm
parents: 27071
diff changeset
   427
  The configuration option @{text "depth_limit"} limits the number of
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   428
  recursive invocations of the simplifier during conditional
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   429
  rewriting.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   430
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   431
  \medskip The Splitter package is usually configured to work as part
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   432
  of the Simplifier.  The effect of repeatedly applying @{ML
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   433
  split_tac} can be simulated by ``@{text "(simp only: split:
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   434
  a\<^sub>1 \<dots> a\<^sub>n)"}''.  There is also a separate @{text split}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   435
  method available for single-step case splitting.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   436
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   437
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   438
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   439
subsection {* Declaring rules *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   440
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   441
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   442
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   443
    @{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   444
    @{attribute_def simp} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   445
    @{attribute_def cong} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   446
    @{attribute_def split} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   447
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   448
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   449
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   450
    (@@{attribute simp} | @@{attribute cong} | @@{attribute split}) (() | 'add' | 'del')
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   451
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   452
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   453
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   454
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   455
  \item @{command "print_simpset"} prints the collection of rules
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   456
  declared to the Simplifier, which is also known as ``simpset''
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   457
  internally \cite{isabelle-ref}.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   458
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   459
  \item @{attribute simp} declares simplification rules.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   460
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   461
  \item @{attribute cong} declares congruence rules.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   462
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   463
  \item @{attribute split} declares case split rules.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   464
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   465
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   466
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   467
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   468
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   469
subsection {* Simplification procedures *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   470
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   471
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   472
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   473
    @{command_def "simproc_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   474
    simproc & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   475
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   476
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   477
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   478
    @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '='
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   479
      @{syntax text} \\ (@'identifier' (@{syntax nameref}+))?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   480
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   481
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   482
    @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+)
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   483
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   484
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   485
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   486
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   487
  \item @{command "simproc_setup"} defines a named simplification
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   488
  procedure that is invoked by the Simplifier whenever any of the
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   489
  given term patterns match the current redex.  The implementation,
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   490
  which is provided as ML source text, needs to be of type @{ML_type
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   491
  "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   492
  cterm} represents the current redex @{text r} and the result is
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   493
  supposed to be some proven rewrite rule @{text "r \<equiv> r'"} (or a
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   494
  generalized version), or @{ML NONE} to indicate failure.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   495
  @{ML_type simpset} argument holds the full context of the current
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   496
  Simplifier invocation, including the actual Isar proof context.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   497
  @{ML_type morphism} informs about the difference of the original
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   498
  compilation context wrt.\ the one of the actual application later
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   499
  on.  The optional @{keyword "identifier"} specifies theorems that
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   500
  represent the logical content of the abstract theory of this
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   501
  simproc.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   502
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   503
  Morphisms and identifiers are only relevant for simprocs that are
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   504
  defined within a local target context, e.g.\ in a locale.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   505
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   506
  \item @{text "simproc add: name"} and @{text "simproc del: name"}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   507
  add or delete named simprocs to the current Simplifier context.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   508
  default is to add a simproc.  Note that @{command "simproc_setup"}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   509
  already adds the new simproc to the subsequent context.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   510
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   511
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   512
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   513
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   514
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   515
subsection {* Forward simplification *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   516
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   517
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   518
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   519
    @{attribute_def simplified} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   520
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   521
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   522
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   523
    @@{attribute simplified} opt? @{syntax thmrefs}?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   524
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   525
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 35613
diff changeset
   526
    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')'
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   527
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   528
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   529
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   530
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   531
  \item @{attribute simplified}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} causes a theorem to
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   532
  be simplified, either by exactly the specified rules @{text "a\<^sub>1, \<dots>,
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   533
  a\<^sub>n"}, or the implicit Simplifier context if no arguments are given.
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   534
  The result is fully simplified by default, including assumptions and
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   535
  conclusion; the options @{text no_asm} etc.\ tune the Simplifier in
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   536
  the same way as the for the @{text simp} method.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   537
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   538
  Note that forward simplification restricts the simplifier to its
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   539
  most basic operation of term rewriting; solver and looper tactics
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   540
  \cite{isabelle-ref} are \emph{not} involved here.  The @{text
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   541
  simplified} attribute should be only rarely required under normal
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   542
  circumstances.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   543
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   544
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   545
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   546
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   547
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   548
section {* The Classical Reasoner \label{sec:classical} *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   549
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   550
subsection {* Basic methods *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   551
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   552
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   553
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   554
    @{method_def rule} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   555
    @{method_def contradiction} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   556
    @{method_def intro} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   557
    @{method_def elim} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   558
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   559
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   560
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   561
    (@@{method rule} | @@{method intro} | @@{method elim}) @{syntax thmrefs}?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   562
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   563
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   564
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   565
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   566
  \item @{method rule} as offered by the Classical Reasoner is a
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   567
  refinement over the primitive one (see \secref{sec:pure-meth-att}).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   568
  Both versions essentially work the same, but the classical version
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   569
  observes the classical rule context in addition to that of
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   570
  Isabelle/Pure.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   571
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   572
  Common object logics (HOL, ZF, etc.) declare a rich collection of
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   573
  classical rules (even if these would qualify as intuitionistic
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   574
  ones), but only few declarations to the rule context of
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   575
  Isabelle/Pure (\secref{sec:pure-meth-att}).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   576
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   577
  \item @{method contradiction} solves some goal by contradiction,
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   578
  deriving any result from both @{text "\<not> A"} and @{text A}.  Chained
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   579
  facts, which are guaranteed to participate, may appear in either
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   580
  order.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   581
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   582
  \item @{method intro} and @{method elim} repeatedly refine some goal
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   583
  by intro- or elim-resolution, after having inserted any chained
26901
d1694ef6e7a7 fixed some Isar element markups;
wenzelm
parents: 26894
diff changeset
   584
  facts.  Exactly the rules given as arguments are taken into account;
d1694ef6e7a7 fixed some Isar element markups;
wenzelm
parents: 26894
diff changeset
   585
  this allows fine-tuned decomposition of a proof problem, in contrast
d1694ef6e7a7 fixed some Isar element markups;
wenzelm
parents: 26894
diff changeset
   586
  to common automated tools.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   587
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   588
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   589
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   590
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   591
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   592
subsection {* Automated methods *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   593
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   594
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   595
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   596
    @{method_def blast} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   597
    @{method_def fast} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   598
    @{method_def slow} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   599
    @{method_def best} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   600
    @{method_def safe} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   601
    @{method_def clarify} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   602
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   603
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   604
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   605
    @@{method blast} @{syntax nat}? (@{syntax clamod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   606
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   607
    (@@{method fast} | @@{method slow} | @@{method best} | @@{method safe} | @@{method clarify})
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   608
      (@{syntax clamod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   609
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   610
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   611
    @{syntax_def clamod}:
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   612
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   613
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   614
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   615
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   616
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   617
  \item @{method blast} refers to the classical tableau prover (see
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   618
  @{ML blast_tac} in \cite{isabelle-ref}).  The optional argument
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   619
  specifies a user-supplied search bound (default 20).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   620
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   621
  \item @{method fast}, @{method slow}, @{method best}, @{method
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   622
  safe}, and @{method clarify} refer to the generic classical
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   623
  reasoner.  See @{ML fast_tac}, @{ML slow_tac}, @{ML best_tac}, @{ML
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   624
  safe_tac}, and @{ML clarify_tac} in \cite{isabelle-ref} for more
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   625
  information.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   626
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   627
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   628
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   629
  Any of the above methods support additional modifiers of the context
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   630
  of classical rules.  Their semantics is analogous to the attributes
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   631
  given before.  Facts provided by forward chaining are inserted into
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 30462
diff changeset
   632
  the goal before commencing proof search.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   633
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   634
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   635
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   636
subsection {* Combined automated methods \label{sec:clasimp} *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   637
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   638
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   639
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   640
    @{method_def auto} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   641
    @{method_def force} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   642
    @{method_def clarsimp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   643
    @{method_def fastsimp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   644
    @{method_def slowsimp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   645
    @{method_def bestsimp} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   646
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   647
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   648
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   649
    @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   650
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   651
    (@@{method force} | @@{method clarsimp} | @@{method fastsimp} | @@{method slowsimp} |
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   652
      @@{method bestsimp}) (@{syntax clasimpmod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   653
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   654
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   655
    @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') |
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   656
      ('cong' | 'split') (() | 'add' | 'del') |
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   657
      'iff' (((() | 'add') '?'?) | 'del') |
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   658
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   659
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   660
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   661
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   662
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   663
  \item @{method auto}, @{method force}, @{method clarsimp}, @{method
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   664
  fastsimp}, @{method slowsimp}, and @{method bestsimp} provide access
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   665
  to Isabelle's combined simplification and classical reasoning
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   666
  tactics.  These correspond to @{ML auto_tac}, @{ML force_tac}, @{ML
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   667
  clarsimp_tac}, and Classical Reasoner tactics with the Simplifier
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   668
  added as wrapper, see \cite{isabelle-ref} for more information.  The
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   669
  modifier arguments correspond to those given in
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   670
  \secref{sec:simplifier} and \secref{sec:classical}.  Just note that
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   671
  the ones related to the Simplifier are prefixed by @{text simp}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   672
  here.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   673
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   674
  Facts provided by forward chaining are inserted into the goal before
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 30462
diff changeset
   675
  doing the search.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   676
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   677
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   678
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   679
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   680
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   681
subsection {* Declaring rules *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   682
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   683
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   684
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   685
    @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   686
    @{attribute_def intro} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   687
    @{attribute_def elim} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   688
    @{attribute_def dest} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   689
    @{attribute_def rule} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   690
    @{attribute_def iff} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   691
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   692
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   693
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   694
    (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   695
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   696
    @@{attribute rule} 'del'
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   697
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   698
    @@{attribute iff} (((() | 'add') '?'?) | 'del')
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   699
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   700
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   701
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   702
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   703
  \item @{command "print_claset"} prints the collection of rules
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   704
  declared to the Classical Reasoner, which is also known as
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   705
  ``claset'' internally \cite{isabelle-ref}.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   706
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   707
  \item @{attribute intro}, @{attribute elim}, and @{attribute dest}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   708
  declare introduction, elimination, and destruction rules,
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   709
  respectively.  By default, rules are considered as \emph{unsafe}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   710
  (i.e.\ not applied blindly without backtracking), while ``@{text
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   711
  "!"}'' classifies as \emph{safe}.  Rule declarations marked by
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   712
  ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   713
  \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   714
  of the @{method rule} method).  The optional natural number
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   715
  specifies an explicit weight argument, which is ignored by automated
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   716
  tools, but determines the search order of single rule steps.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   717
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   718
  \item @{attribute rule}~@{text del} deletes introduction,
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   719
  elimination, or destruction rules from the context.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   720
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   721
  \item @{attribute iff} declares logical equivalences to the
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   722
  Simplifier and the Classical reasoner at the same time.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   723
  Non-conditional rules result in a ``safe'' introduction and
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   724
  elimination pair; conditional ones are considered ``unsafe''.  Rules
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   725
  with negative conclusion are automatically inverted (using @{text
26789
fc6d5fa0ca3c misc fixes and tuning;
wenzelm
parents: 26782
diff changeset
   726
  "\<not>"}-elimination internally).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   727
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   728
  The ``@{text "?"}'' version of @{attribute iff} declares rules to
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   729
  the Isabelle/Pure context only, and omits the Simplifier
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   730
  declaration.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   731
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   732
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   733
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   734
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   735
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   736
subsection {* Classical operations *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   737
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   738
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   739
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   740
    @{attribute_def swapped} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   741
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   742
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   743
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   744
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   745
  \item @{attribute swapped} turns an introduction rule into an
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   746
  elimination, by resolving with the classical swap principle @{text
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   747
  "(\<not> B \<Longrightarrow> A) \<Longrightarrow> (\<not> A \<Longrightarrow> B)"}.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   748
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   749
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   750
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   751
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   752
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   753
section {* Object-logic setup \label{sec:object-logic} *}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   754
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   755
text {*
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   756
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   757
    @{command_def "judgment"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   758
    @{method_def atomize} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   759
    @{attribute_def atomize} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   760
    @{attribute_def rule_format} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   761
    @{attribute_def rulify} & : & @{text attribute} \\
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   762
  \end{matharray}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   763
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   764
  The very starting point for any Isabelle object-logic is a ``truth
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   765
  judgment'' that links object-level statements to the meta-logic
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   766
  (with its minimal language of @{text prop} that covers universal
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   767
  quantification @{text "\<And>"} and implication @{text "\<Longrightarrow>"}).
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   768
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   769
  Common object-logics are sufficiently expressive to internalize rule
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   770
  statements over @{text "\<And>"} and @{text "\<Longrightarrow>"} within their own
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   771
  language.  This is useful in certain situations where a rule needs
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   772
  to be viewed as an atomic statement from the meta-level perspective,
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   773
  e.g.\ @{text "\<And>x. x \<in> A \<Longrightarrow> P x"} versus @{text "\<forall>x \<in> A. P x"}.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   774
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   775
  From the following language elements, only the @{method atomize}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   776
  method and @{attribute rule_format} attribute are occasionally
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   777
  required by end-users, the rest is for those who need to setup their
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   778
  own object-logic.  In the latter case existing formulations of
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   779
  Isabelle/FOL or Isabelle/HOL may be taken as realistic examples.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   780
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   781
  Generic tools may refer to the information provided by object-logic
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   782
  declarations internally.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   783
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   784
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   785
    @@{command judgment} @{syntax constdecl}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   786
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   787
    @@{attribute atomize} ('(' 'full' ')')?
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   788
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   789
    @@{attribute rule_format} ('(' 'noasm' ')')?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   790
  "}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   791
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   792
  \begin{description}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   793
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   794
  \item @{command "judgment"}~@{text "c :: \<sigma> (mx)"} declares constant
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   795
  @{text c} as the truth judgment of the current object-logic.  Its
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   796
  type @{text \<sigma>} should specify a coercion of the category of
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   797
  object-level propositions to @{text prop} of the Pure meta-logic;
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   798
  the mixfix annotation @{text "(mx)"} would typically just link the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   799
  object language (internally of syntactic category @{text logic})
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   800
  with that of @{text prop}.  Only one @{command "judgment"}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   801
  declaration may be given in any theory development.
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   802
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   803
  \item @{method atomize} (as a method) rewrites any non-atomic
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   804
  premises of a sub-goal, using the meta-level equations declared via
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   805
  @{attribute atomize} (as an attribute) beforehand.  As a result,
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   806
  heavily nested goals become amenable to fundamental operations such
42626
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 42617
diff changeset
   807
  as resolution (cf.\ the @{method (Pure) rule} method).  Giving the ``@{text
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   808
  "(full)"}'' option here means to turn the whole subgoal into an
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   809
  object-statement (if possible), including the outermost parameters
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   810
  and assumptions as well.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   811
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   812
  A typical collection of @{attribute atomize} rules for a particular
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   813
  object-logic would provide an internalization for each of the
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   814
  connectives of @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"}.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   815
  Meta-level conjunction should be covered as well (this is
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   816
  particularly important for locales, see \secref{sec:locale}).
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   817
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   818
  \item @{attribute rule_format} rewrites a theorem by the equalities
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   819
  declared as @{attribute rulify} rules in the current object-logic.
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   820
  By default, the result is fully normalized, including assumptions
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   821
  and conclusions at any depth.  The @{text "(no_asm)"} option
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   822
  restricts the transformation to the conclusion of a rule.
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   823
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   824
  In common object-logics (HOL, FOL, ZF), the effect of @{attribute
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   825
  rule_format} is to replace (bounded) universal quantification
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   826
  (@{text "\<forall>"}) and implication (@{text "\<longrightarrow>"}) by the corresponding
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   827
  rule statements over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   828
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   829
  \end{description}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   830
*}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
   831
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   832
end