doc-src/IsarRef/Thy/Generic.thy
author wenzelm
Thu, 09 Jun 2011 22:25:25 +0200
changeset 43332 dca2c7c598f0
parent 42930 41394a61cca9
child 43365 f8944cb2468f
permissions -rw-r--r--
document depth arguments of method "auto";
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
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   471
text {* Simplification procedures are ML functions that produce proven
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   472
  rewrite rules on demand.  They are associated with higher-order
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   473
  patterns that approximate the left-hand sides of equations.  The
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   474
  Simplifier first matches the current redex against one of the LHS
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   475
  patterns; if this succeeds, the corresponding ML function is
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   476
  invoked, passing the Simplifier context and redex term.  Thus rules
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   477
  may be specifically fashioned for particular situations, resulting
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   478
  in a more powerful mechanism than term rewriting by a fixed set of
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   479
  rules.
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   480
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   481
  Any successful result needs to be a (possibly conditional) rewrite
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   482
  rule @{text "t \<equiv> u"} that is applicable to the current redex.  The
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   483
  rule will be applied just as any ordinary rewrite rule.  It is
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   484
  expected to be already in \emph{internal form}, bypassing the
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   485
  automatic preprocessing of object-level equivalences.
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   486
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   487
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   488
    @{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
   489
    simproc & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   490
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   491
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   492
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   493
    @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '='
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   494
      @{syntax text} \\ (@'identifier' (@{syntax nameref}+))?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   495
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   496
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   497
    @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+)
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   498
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   499
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   500
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   501
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   502
  \item @{command "simproc_setup"} defines a named simplification
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   503
  procedure that is invoked by the Simplifier whenever any of the
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   504
  given term patterns match the current redex.  The implementation,
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   505
  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
   506
  "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   507
  cterm} represents the current redex @{text r} and the result is
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   508
  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
   509
  generalized version), or @{ML NONE} to indicate failure.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   510
  @{ML_type simpset} argument holds the full context of the current
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   511
  Simplifier invocation, including the actual Isar proof context.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   512
  @{ML_type morphism} informs about the difference of the original
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   513
  compilation context wrt.\ the one of the actual application later
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   514
  on.  The optional @{keyword "identifier"} specifies theorems that
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   515
  represent the logical content of the abstract theory of this
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   516
  simproc.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   517
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   518
  Morphisms and identifiers are only relevant for simprocs that are
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   519
  defined within a local target context, e.g.\ in a locale.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   520
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   521
  \item @{text "simproc add: name"} and @{text "simproc del: name"}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   522
  add or delete named simprocs to the current Simplifier context.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   523
  default is to add a simproc.  Note that @{command "simproc_setup"}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   524
  already adds the new simproc to the subsequent context.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   525
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   526
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   527
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   528
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   529
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   530
subsubsection {* Example *}
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   531
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   532
text {* The following simplification procedure for @{thm
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   533
  [source=false, show_types] unit_eq} in HOL performs fine-grained
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   534
  control over rule application, beyond higher-order pattern matching.
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   535
  Declaring @{thm unit_eq} as @{attribute simp} directly would make
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   536
  the simplifier loop!  Note that a version of this simplification
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   537
  procedure is already active in Isabelle/HOL.  *}
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   538
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   539
simproc_setup unit ("x::unit") = {*
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   540
  fn _ => fn _ => fn ct =>
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   541
    if HOLogic.is_unit (term_of ct) then NONE
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   542
    else SOME (mk_meta_eq @{thm unit_eq})
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   543
*}
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   544
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   545
text {* Since the Simplifier applies simplification procedures
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   546
  frequently, it is important to make the failure check in ML
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   547
  reasonably fast. *}
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   548
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   549
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   550
subsection {* Forward simplification *}
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
    @{attribute_def simplified} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   555
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   556
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   557
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   558
    @@{attribute simplified} opt? @{syntax thmrefs}?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   559
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   560
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 35613
diff changeset
   561
    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')'
42596
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 @{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
   567
  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
   568
  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
   569
  The result is fully simplified by default, including assumptions and
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   570
  conclusion; the options @{text no_asm} etc.\ tune the Simplifier in
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   571
  the same way as the for the @{text simp} method.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   572
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   573
  Note that forward simplification restricts the simplifier to its
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   574
  most basic operation of term rewriting; solver and looper tactics
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   575
  \cite{isabelle-ref} are \emph{not} involved here.  The @{text
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   576
  simplified} attribute should be only rarely required under normal
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   577
  circumstances.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   578
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   579
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   580
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   581
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   582
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   583
section {* The Classical Reasoner \label{sec:classical} *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   584
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   585
subsection {* Basic concepts *}
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   586
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   587
text {* Although Isabelle is generic, many users will be working in
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   588
  some extension of classical first-order logic.  Isabelle/ZF is built
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   589
  upon theory FOL, while Isabelle/HOL conceptually contains
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   590
  first-order logic as a fragment.  Theorem-proving in predicate logic
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   591
  is undecidable, but many automated strategies have been developed to
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   592
  assist in this task.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   593
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   594
  Isabelle's classical reasoner is a generic package that accepts
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   595
  certain information about a logic and delivers a suite of automatic
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   596
  proof tools, based on rules that are classified and declared in the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   597
  context.  These proof procedures are slow and simplistic compared
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   598
  with high-end automated theorem provers, but they can save
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   599
  considerable time and effort in practice.  They can prove theorems
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   600
  such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   601
  milliseconds (including full proof reconstruction): *}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   602
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   603
lemma "(\<exists>y. \<forall>x. F x y \<longleftrightarrow> F x x) \<longrightarrow> \<not> (\<forall>x. \<exists>y. \<forall>z. F z y \<longleftrightarrow> \<not> F z x)"
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   604
  by blast
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   605
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   606
lemma "(\<forall>z. \<exists>y. \<forall>x. f x y \<longleftrightarrow> f x z \<and> \<not> f x x) \<longrightarrow> \<not> (\<exists>z. \<forall>x. f x z)"
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   607
  by blast
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   608
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   609
text {* The proof tools are generic.  They are not restricted to
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   610
  first-order logic, and have been heavily used in the development of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   611
  the Isabelle/HOL library and applications.  The tactics can be
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   612
  traced, and their components can be called directly; in this manner,
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   613
  any proof can be viewed interactively.  *}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   614
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   615
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   616
subsubsection {* The sequent calculus *}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   617
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   618
text {* Isabelle supports natural deduction, which is easy to use for
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   619
  interactive proof.  But natural deduction does not easily lend
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   620
  itself to automation, and has a bias towards intuitionism.  For
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   621
  certain proofs in classical logic, it can not be called natural.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   622
  The \emph{sequent calculus}, a generalization of natural deduction,
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   623
  is easier to automate.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   624
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   625
  A \textbf{sequent} has the form @{text "\<Gamma> \<turnstile> \<Delta>"}, where @{text "\<Gamma>"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   626
  and @{text "\<Delta>"} are sets of formulae.\footnote{For first-order
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   627
  logic, sequents can equivalently be made from lists or multisets of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   628
  formulae.} The sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} is
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   629
  \textbf{valid} if @{text "P\<^sub>1 \<and> \<dots> \<and> P\<^sub>m"} implies @{text "Q\<^sub>1 \<or> \<dots> \<or>
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   630
  Q\<^sub>n"}.  Thus @{text "P\<^sub>1, \<dots>, P\<^sub>m"} represent assumptions, each of which
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   631
  is true, while @{text "Q\<^sub>1, \<dots>, Q\<^sub>n"} represent alternative goals.  A
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   632
  sequent is \textbf{basic} if its left and right sides have a common
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   633
  formula, as in @{text "P, Q \<turnstile> Q, R"}; basic sequents are trivially
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   634
  valid.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   635
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   636
  Sequent rules are classified as \textbf{right} or \textbf{left},
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   637
  indicating which side of the @{text "\<turnstile>"} symbol they operate on.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   638
  Rules that operate on the right side are analogous to natural
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   639
  deduction's introduction rules, and left rules are analogous to
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   640
  elimination rules.  The sequent calculus analogue of @{text "(\<longrightarrow>I)"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   641
  is the rule
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   642
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   643
  \infer[@{text "(\<longrightarrow>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<longrightarrow> Q"}}{@{text "P, \<Gamma> \<turnstile> \<Delta>, Q"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   644
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   645
  Applying the rule backwards, this breaks down some implication on
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   646
  the right side of a sequent; @{text "\<Gamma>"} and @{text "\<Delta>"} stand for
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   647
  the sets of formulae that are unaffected by the inference.  The
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   648
  analogue of the pair @{text "(\<or>I1)"} and @{text "(\<or>I2)"} is the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   649
  single rule
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   650
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   651
  \infer[@{text "(\<or>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<or> Q"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P, Q"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   652
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   653
  This breaks down some disjunction on the right side, replacing it by
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   654
  both disjuncts.  Thus, the sequent calculus is a kind of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   655
  multiple-conclusion logic.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   656
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   657
  To illustrate the use of multiple formulae on the right, let us
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   658
  prove the classical theorem @{text "(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}.  Working
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   659
  backwards, we reduce this formula to a basic sequent:
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   660
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   661
  \infer[@{text "(\<or>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   662
    {\infer[@{text "(\<longrightarrow>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q), (Q \<longrightarrow> P)"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   663
      {\infer[@{text "(\<longrightarrow>R)"}]{@{text "P \<turnstile> Q, (Q \<longrightarrow> P)"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   664
        {@{text "P, Q \<turnstile> Q, P"}}}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   665
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   666
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   667
  This example is typical of the sequent calculus: start with the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   668
  desired theorem and apply rules backwards in a fairly arbitrary
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   669
  manner.  This yields a surprisingly effective proof procedure.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   670
  Quantifiers add only few complications, since Isabelle handles
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   671
  parameters and schematic variables.  See \cite[Chapter
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   672
  10]{paulson-ml2} for further discussion.  *}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   673
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   674
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   675
subsubsection {* Simulating sequents by natural deduction *}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   676
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   677
text {* Isabelle can represent sequents directly, as in the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   678
  object-logic LK.  But natural deduction is easier to work with, and
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   679
  most object-logics employ it.  Fortunately, we can simulate the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   680
  sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} by the Isabelle formula
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   681
  @{text "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>2 \<Longrightarrow> ... \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> Q\<^sub>1"} where the order of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   682
  the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   683
  Elim-resolution plays a key role in simulating sequent proofs.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   684
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   685
  We can easily handle reasoning on the left.  Elim-resolution with
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   686
  the rules @{text "(\<or>E)"}, @{text "(\<bottom>E)"} and @{text "(\<exists>E)"} achieves
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   687
  a similar effect as the corresponding sequent rules.  For the other
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   688
  connectives, we use sequent-style elimination rules instead of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   689
  destruction rules such as @{text "(\<and>E1, 2)"} and @{text "(\<forall>E)"}.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   690
  But note that the rule @{text "(\<not>L)"} has no effect under our
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   691
  representation of sequents!
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   692
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   693
  \infer[@{text "(\<not>L)"}]{@{text "\<not> P, \<Gamma> \<turnstile> \<Delta>"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   694
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   695
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   696
  What about reasoning on the right?  Introduction rules can only
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   697
  affect the formula in the conclusion, namely @{text "Q\<^sub>1"}.  The
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   698
  other right-side formulae are represented as negated assumptions,
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   699
  @{text "\<not> Q\<^sub>2, \<dots>, \<not> Q\<^sub>n"}.  In order to operate on one of these, it
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   700
  must first be exchanged with @{text "Q\<^sub>1"}.  Elim-resolution with the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   701
  @{text swap} rule has this effect: @{text "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   702
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   703
  To ensure that swaps occur only when necessary, each introduction
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   704
  rule is converted into a swapped form: it is resolved with the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   705
  second premise of @{text "(swap)"}.  The swapped form of @{text
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   706
  "(\<and>I)"}, which might be called @{text "(\<not>\<and>E)"}, is
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   707
  @{text [display] "\<not> (P \<and> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (\<not> R \<Longrightarrow> Q) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   708
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   709
  Similarly, the swapped form of @{text "(\<longrightarrow>I)"} is
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   710
  @{text [display] "\<not> (P \<longrightarrow> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P \<Longrightarrow> Q) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   711
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   712
  Swapped introduction rules are applied using elim-resolution, which
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   713
  deletes the negated formula.  Our representation of sequents also
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   714
  requires the use of ordinary introduction rules.  If we had no
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   715
  regard for readability of intermediate goal states, we could treat
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   716
  the right side more uniformly by representing sequents as @{text
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   717
  [display] "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> \<bottom>"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   718
*}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   719
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   720
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   721
subsubsection {* Extra rules for the sequent calculus *}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   722
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   723
text {* As mentioned, destruction rules such as @{text "(\<and>E1, 2)"} and
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   724
  @{text "(\<forall>E)"} must be replaced by sequent-style elimination rules.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   725
  In addition, we need rules to embody the classical equivalence
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   726
  between @{text "P \<longrightarrow> Q"} and @{text "\<not> P \<or> Q"}.  The introduction
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   727
  rules @{text "(\<or>I1, 2)"} are replaced by a rule that simulates
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   728
  @{text "(\<or>R)"}: @{text [display] "(\<not> Q \<Longrightarrow> P) \<Longrightarrow> P \<or> Q"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   729
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   730
  The destruction rule @{text "(\<longrightarrow>E)"} is replaced by @{text [display]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   731
  "(P \<longrightarrow> Q) \<Longrightarrow> (\<not> P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   732
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   733
  Quantifier replication also requires special rules.  In classical
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   734
  logic, @{text "\<exists>x. P x"} is equivalent to @{text "\<not> (\<forall>x. \<not> P x)"};
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   735
  the rules @{text "(\<exists>R)"} and @{text "(\<forall>L)"} are dual:
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   736
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   737
  \infer[@{text "(\<exists>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x"}}{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x, P t"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   738
  \qquad
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   739
  \infer[@{text "(\<forall>L)"}]{@{text "\<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}{@{text "P t, \<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   740
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   741
  Thus both kinds of quantifier may be replicated.  Theorems requiring
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   742
  multiple uses of a universal formula are easy to invent; consider
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   743
  @{text [display] "(\<forall>x. P x \<longrightarrow> P (f x)) \<and> P a \<longrightarrow> P (f\<^sup>n a)"} for any
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   744
  @{text "n > 1"}.  Natural examples of the multiple use of an
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   745
  existential formula are rare; a standard one is @{text "\<exists>x. \<forall>y. P x
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   746
  \<longrightarrow> P y"}.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   747
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   748
  Forgoing quantifier replication loses completeness, but gains
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   749
  decidability, since the search space becomes finite.  Many useful
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   750
  theorems can be proved without replication, and the search generally
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   751
  delivers its verdict in a reasonable time.  To adopt this approach,
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   752
  represent the sequent rules @{text "(\<exists>R)"}, @{text "(\<exists>L)"} and
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   753
  @{text "(\<forall>R)"} by @{text "(\<exists>I)"}, @{text "(\<exists>E)"} and @{text "(\<forall>I)"},
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   754
  respectively, and put @{text "(\<forall>E)"} into elimination form: @{text
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   755
  [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   756
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   757
  Elim-resolution with this rule will delete the universal formula
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   758
  after a single use.  To replicate universal quantifiers, replace the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   759
  rule by @{text [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   760
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   761
  To replicate existential quantifiers, replace @{text "(\<exists>I)"} by
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   762
  @{text [display] "(\<not> (\<exists>x. P x) \<Longrightarrow> P t) \<Longrightarrow> \<exists>x. P x"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   763
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   764
  All introduction rules mentioned above are also useful in swapped
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   765
  form.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   766
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   767
  Replication makes the search space infinite; we must apply the rules
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   768
  with care.  The classical reasoner distinguishes between safe and
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   769
  unsafe rules, applying the latter only when there is no alternative.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   770
  Depth-first search may well go down a blind alley; best-first search
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   771
  is better behaved in an infinite search space.  However, quantifier
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   772
  replication is too expensive to prove any but the simplest theorems.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   773
*}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   774
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
   775
42928
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   776
subsection {* Rule declarations *}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   777
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   778
text {* The proof tools of the Classical Reasoner depend on
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   779
  collections of rules declared in the context, which are classified
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   780
  as introduction, elimination or destruction and as \emph{safe} or
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   781
  \emph{unsafe}.  In general, safe rules can be attempted blindly,
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   782
  while unsafe rules must be used with care.  A safe rule must never
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   783
  reduce a provable goal to an unprovable set of subgoals.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   784
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   785
  The rule @{text "P \<Longrightarrow> P \<or> Q"} is unsafe because it reduces @{text "P
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   786
  \<or> Q"} to @{text "P"}, which might turn out as premature choice of an
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   787
  unprovable subgoal.  Any rule is unsafe whose premises contain new
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   788
  unknowns.  The elimination rule @{text "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"} is
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   789
  unsafe, since it is applied via elim-resolution, which discards the
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   790
  assumption @{text "\<forall>x. P x"} and replaces it by the weaker
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   791
  assumption @{text "P t"}.  The rule @{text "P t \<Longrightarrow> \<exists>x. P x"} is
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   792
  unsafe for similar reasons.  The quantifier duplication rule @{text
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   793
  "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"} is unsafe in a different sense:
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   794
  since it keeps the assumption @{text "\<forall>x. P x"}, it is prone to
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   795
  looping.  In classical first-order logic, all rules are safe except
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   796
  those mentioned above.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   797
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   798
  The safe~/ unsafe distinction is vague, and may be regarded merely
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   799
  as a way of giving some rules priority over others.  One could argue
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   800
  that @{text "(\<or>E)"} is unsafe, because repeated application of it
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   801
  could generate exponentially many subgoals.  Induction rules are
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   802
  unsafe because inductive proofs are difficult to set up
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   803
  automatically.  Any inference is unsafe that instantiates an unknown
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   804
  in the proof state --- thus matching must be used, rather than
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   805
  unification.  Even proof by assumption is unsafe if it instantiates
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   806
  unknowns shared with other subgoals.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   807
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   808
  \begin{matharray}{rcl}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   809
    @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   810
    @{attribute_def intro} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   811
    @{attribute_def elim} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   812
    @{attribute_def dest} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   813
    @{attribute_def rule} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   814
    @{attribute_def iff} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   815
    @{attribute_def swapped} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   816
  \end{matharray}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   817
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   818
  @{rail "
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   819
    (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}?
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   820
    ;
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   821
    @@{attribute rule} 'del'
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   822
    ;
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   823
    @@{attribute iff} (((() | 'add') '?'?) | 'del')
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   824
  "}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   825
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   826
  \begin{description}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   827
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   828
  \item @{command "print_claset"} prints the collection of rules
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   829
  declared to the Classical Reasoner, i.e.\ the @{ML_type claset}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   830
  within the context.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   831
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   832
  \item @{attribute intro}, @{attribute elim}, and @{attribute dest}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   833
  declare introduction, elimination, and destruction rules,
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   834
  respectively.  By default, rules are considered as \emph{unsafe}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   835
  (i.e.\ not applied blindly without backtracking), while ``@{text
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   836
  "!"}'' classifies as \emph{safe}.  Rule declarations marked by
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   837
  ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   838
  \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   839
  of the @{method rule} method).  The optional natural number
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   840
  specifies an explicit weight argument, which is ignored by the
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   841
  automated reasoning tools, but determines the search order of single
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   842
  rule steps.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   843
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   844
  Introduction rules are those that can be applied using ordinary
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   845
  resolution.  Their swapped forms are generated internally, which
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   846
  will be applied using elim-resolution.  Elimination rules are
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   847
  applied using elim-resolution.  Rules are sorted by the number of
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   848
  new subgoals they will yield; rules that generate the fewest
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   849
  subgoals will be tried first.  Otherwise, later declarations take
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   850
  precedence over earlier ones.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   851
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   852
  Rules already present in the context with the same classification
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   853
  are ignored.  A warning is printed if the rule has already been
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   854
  added with some other classification, but the rule is added anyway
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   855
  as requested.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   856
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   857
  \item @{attribute rule}~@{text del} deletes all occurrences of a
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   858
  rule from the classical context, regardless of its classification as
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   859
  introduction~/ elimination~/ destruction and safe~/ unsafe.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   860
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   861
  \item @{attribute iff} declares logical equivalences to the
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   862
  Simplifier and the Classical reasoner at the same time.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   863
  Non-conditional rules result in a safe introduction and elimination
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   864
  pair; conditional ones are considered unsafe.  Rules with negative
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   865
  conclusion are automatically inverted (using @{text "\<not>"}-elimination
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   866
  internally).
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   867
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   868
  The ``@{text "?"}'' version of @{attribute iff} declares rules to
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   869
  the Isabelle/Pure context only, and omits the Simplifier
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   870
  declaration.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   871
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   872
  \item @{attribute swapped} turns an introduction rule into an
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   873
  elimination, by resolving with the classical swap principle @{text
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   874
  "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"} in the second position.  This is mainly for
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   875
  illustrative purposes: the Classical Reasoner already swaps rules
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   876
  internally as explained above.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
   877
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   878
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   879
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   880
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   881
27040
3d3e6e07b931 major reorganization of document structure;
wenzelm
parents: 26901
diff changeset
   882
subsection {* Automated methods *}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   883
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   884
text {*
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   885
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   886
    @{method_def blast} & : & @{text method} \\
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   887
    @{method_def auto} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   888
    @{method_def force} & : & @{text method} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   889
    @{method_def fast} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   890
    @{method_def slow} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   891
    @{method_def best} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   892
    @{method_def fastsimp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   893
    @{method_def slowsimp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   894
    @{method_def bestsimp} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   895
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   896
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   897
  @{rail "
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   898
    @@{method blast} @{syntax nat}? (@{syntax clamod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   899
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   900
    @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   901
    ;
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   902
    @@{method force} (@{syntax clasimpmod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   903
    ;
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   904
    (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   905
    ;
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   906
    (@@{method fastsimp} | @@{method slowsimp} | @@{method bestsimp})
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   907
      (@{syntax clasimpmod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   908
    ;
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   909
    @{syntax_def clamod}:
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   910
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   911
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   912
    @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') |
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   913
      ('cong' | 'split') (() | 'add' | 'del') |
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   914
      'iff' (((() | 'add') '?'?) | 'del') |
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   915
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs}
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   916
  "}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   917
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   918
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   919
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   920
  \item @{method blast} is a separate classical tableau prover that
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   921
  uses the same classical rule declarations as explained before.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   922
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   923
  Proof search is coded directly in ML using special data structures.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   924
  A successful proof is then reconstructed using regular Isabelle
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   925
  inferences.  It is faster and more powerful than the other classical
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   926
  reasoning tools, but has major limitations too.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   927
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   928
  \begin{itemize}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   929
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   930
  \item It does not use the classical wrapper tacticals, such as the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   931
  integration with the Simplifier of @{method fastsimp}.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   932
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   933
  \item It does not perform higher-order unification, as needed by the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   934
  rule @{thm [source=false] rangeI} in HOL.  There are often
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   935
  alternatives to such rules, for example @{thm [source=false]
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   936
  range_eqI}.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   937
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   938
  \item Function variables may only be applied to parameters of the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   939
  subgoal.  (This restriction arises because the prover does not use
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   940
  higher-order unification.)  If other function variables are present
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   941
  then the prover will fail with the message \texttt{Function Var's
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   942
  argument not a bound variable}.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   943
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   944
  \item Its proof strategy is more general than @{method fast} but can
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   945
  be slower.  If @{method blast} fails or seems to be running forever,
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   946
  try @{method fast} and the other proof tools described below.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   947
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   948
  \end{itemize}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   949
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   950
  The optional integer argument specifies a bound for the number of
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   951
  unsafe steps used in a proof.  By default, @{method blast} starts
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   952
  with a bound of 0 and increases it successively to 20.  In contrast,
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   953
  @{text "(blast lim)"} tries to prove the goal using a search bound
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   954
  of @{text "lim"}.  Sometimes a slow proof using @{method blast} can
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   955
  be made much faster by supplying the successful search bound to this
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   956
  proof method instead.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   957
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   958
  \item @{method auto} combines classical reasoning with
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   959
  simplification.  It is intended for situations where there are a lot
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   960
  of mostly trivial subgoals; it proves all the easy ones, leaving the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   961
  ones it cannot prove.  Occasionally, attempting to prove the hard
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   962
  ones may take a long time.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   963
43332
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
   964
  The optional depth arguments in @{text "(auto m n)"} refer to its
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
   965
  builtin classical reasoning procedures: @{text m} (default 4) is for
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
   966
  @{method blast}, which is tried first, and @{text n} (default 2) is
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
   967
  for a slower but more general alternative that also takes wrappers
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
   968
  into account.
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   969
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   970
  \item @{method force} is intended to prove the first subgoal
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   971
  completely, using many fancy proof tools and performing a rather
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   972
  exhaustive search.  As a result, proof attempts may take rather long
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   973
  or diverge easily.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   974
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   975
  \item @{method fast}, @{method best}, @{method slow} attempt to
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   976
  prove the first subgoal using sequent-style reasoning as explained
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   977
  before.  Unlike @{method blast}, they construct proofs directly in
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   978
  Isabelle.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   979
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   980
  There is a difference in search strategy and back-tracking: @{method
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   981
  fast} uses depth-first search and @{method best} uses best-first
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   982
  search (guided by a heuristic function: normally the total size of
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   983
  the proof state).
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   984
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   985
  Method @{method slow} is like @{method fast}, but conducts a broader
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   986
  search: it may, when backtracking from a failed proof attempt, undo
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   987
  even the step of proving a subgoal by assumption.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   988
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   989
  \item @{method fastsimp}, @{method slowsimp}, @{method bestsimp} are
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   990
  like @{method fast}, @{method slow}, @{method best}, respectively,
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   991
  but use the Simplifier as additional wrapper.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   992
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   993
  \end{description}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   994
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   995
  Any of the above methods support additional modifiers of the context
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   996
  of classical (and simplifier) rules, but the ones related to the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   997
  Simplifier are explicitly prefixed by @{text simp} here.  The
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   998
  semantics of these ad-hoc rule declarations is analogous to the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
   999
  attributes given before.  Facts provided by forward chaining are
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1000
  inserted into the goal before commencing proof search.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1001
*}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1002
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1003
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1004
subsection {* Semi-automated methods *}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1005
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1006
text {* These proof methods may help in situations when the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1007
  fully-automated tools fail.  The result is a simpler subgoal that
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1008
  can be tackled by other means, such as by manual instantiation of
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1009
  quantifiers.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1010
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1011
  \begin{matharray}{rcl}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1012
    @{method_def safe} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1013
    @{method_def clarify} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1014
    @{method_def clarsimp} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1015
  \end{matharray}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1016
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1017
  @{rail "
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1018
    (@@{method safe} | @@{method clarify}) (@{syntax clamod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1019
    ;
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1020
    @@{method clarsimp} (@{syntax clasimpmod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1021
  "}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1022
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1023
  \begin{description}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1024
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1025
  \item @{method safe} repeatedly performs safe steps on all subgoals.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1026
  It is deterministic, with at most one outcome.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1027
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1028
  \item @{method clarify} performs a series of safe steps as follows.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1029
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1030
  No splitting step is applied; for example, the subgoal @{text "A \<and>
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1031
  B"} is left as a conjunction.  Proof by assumption, Modus Ponens,
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1032
  etc., may be performed provided they do not instantiate unknowns.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1033
  Assumptions of the form @{text "x = t"} may be eliminated.  The safe
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1034
  wrapper tactical is applied.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1035
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1036
  \item @{method clarsimp} acts like @{method clarify}, but also does
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1037
  simplification.  Note that if the Simplifier context includes a
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1038
  splitter for the premises, the subgoal may still be split.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1039
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1040
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1041
*}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1042
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1043
42929
wenzelm
parents: 42928
diff changeset
  1044
subsection {* Structured proof methods *}
wenzelm
parents: 42928
diff changeset
  1045
wenzelm
parents: 42928
diff changeset
  1046
text {*
wenzelm
parents: 42928
diff changeset
  1047
  \begin{matharray}{rcl}
wenzelm
parents: 42928
diff changeset
  1048
    @{method_def rule} & : & @{text method} \\
wenzelm
parents: 42928
diff changeset
  1049
    @{method_def contradiction} & : & @{text method} \\
wenzelm
parents: 42928
diff changeset
  1050
    @{method_def intro} & : & @{text method} \\
wenzelm
parents: 42928
diff changeset
  1051
    @{method_def elim} & : & @{text method} \\
wenzelm
parents: 42928
diff changeset
  1052
  \end{matharray}
wenzelm
parents: 42928
diff changeset
  1053
wenzelm
parents: 42928
diff changeset
  1054
  @{rail "
wenzelm
parents: 42928
diff changeset
  1055
    (@@{method rule} | @@{method intro} | @@{method elim}) @{syntax thmrefs}?
wenzelm
parents: 42928
diff changeset
  1056
  "}
wenzelm
parents: 42928
diff changeset
  1057
wenzelm
parents: 42928
diff changeset
  1058
  \begin{description}
wenzelm
parents: 42928
diff changeset
  1059
wenzelm
parents: 42928
diff changeset
  1060
  \item @{method rule} as offered by the Classical Reasoner is a
wenzelm
parents: 42928
diff changeset
  1061
  refinement over the Pure one (see \secref{sec:pure-meth-att}).  Both
wenzelm
parents: 42928
diff changeset
  1062
  versions work the same, but the classical version observes the
wenzelm
parents: 42928
diff changeset
  1063
  classical rule context in addition to that of Isabelle/Pure.
wenzelm
parents: 42928
diff changeset
  1064
wenzelm
parents: 42928
diff changeset
  1065
  Common object logics (HOL, ZF, etc.) declare a rich collection of
wenzelm
parents: 42928
diff changeset
  1066
  classical rules (even if these would qualify as intuitionistic
wenzelm
parents: 42928
diff changeset
  1067
  ones), but only few declarations to the rule context of
wenzelm
parents: 42928
diff changeset
  1068
  Isabelle/Pure (\secref{sec:pure-meth-att}).
wenzelm
parents: 42928
diff changeset
  1069
wenzelm
parents: 42928
diff changeset
  1070
  \item @{method contradiction} solves some goal by contradiction,
wenzelm
parents: 42928
diff changeset
  1071
  deriving any result from both @{text "\<not> A"} and @{text A}.  Chained
wenzelm
parents: 42928
diff changeset
  1072
  facts, which are guaranteed to participate, may appear in either
wenzelm
parents: 42928
diff changeset
  1073
  order.
wenzelm
parents: 42928
diff changeset
  1074
wenzelm
parents: 42928
diff changeset
  1075
  \item @{method intro} and @{method elim} repeatedly refine some goal
wenzelm
parents: 42928
diff changeset
  1076
  by intro- or elim-resolution, after having inserted any chained
wenzelm
parents: 42928
diff changeset
  1077
  facts.  Exactly the rules given as arguments are taken into account;
wenzelm
parents: 42928
diff changeset
  1078
  this allows fine-tuned decomposition of a proof problem, in contrast
wenzelm
parents: 42928
diff changeset
  1079
  to common automated tools.
wenzelm
parents: 42928
diff changeset
  1080
wenzelm
parents: 42928
diff changeset
  1081
  \end{description}
wenzelm
parents: 42928
diff changeset
  1082
*}
wenzelm
parents: 42928
diff changeset
  1083
wenzelm
parents: 42928
diff changeset
  1084
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
  1085
section {* Object-logic setup \label{sec:object-logic} *}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1086
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1087
text {*
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1088
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1089
    @{command_def "judgment"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1090
    @{method_def atomize} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1091
    @{attribute_def atomize} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1092
    @{attribute_def rule_format} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1093
    @{attribute_def rulify} & : & @{text attribute} \\
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1094
  \end{matharray}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1095
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1096
  The very starting point for any Isabelle object-logic is a ``truth
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1097
  judgment'' that links object-level statements to the meta-logic
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1098
  (with its minimal language of @{text prop} that covers universal
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1099
  quantification @{text "\<And>"} and implication @{text "\<Longrightarrow>"}).
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1100
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1101
  Common object-logics are sufficiently expressive to internalize rule
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1102
  statements over @{text "\<And>"} and @{text "\<Longrightarrow>"} within their own
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1103
  language.  This is useful in certain situations where a rule needs
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1104
  to be viewed as an atomic statement from the meta-level perspective,
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1105
  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
  1106
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1107
  From the following language elements, only the @{method atomize}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1108
  method and @{attribute rule_format} attribute are occasionally
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1109
  required by end-users, the rest is for those who need to setup their
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1110
  own object-logic.  In the latter case existing formulations of
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1111
  Isabelle/FOL or Isabelle/HOL may be taken as realistic examples.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1112
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1113
  Generic tools may refer to the information provided by object-logic
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1114
  declarations internally.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1115
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1116
  @{rail "
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1117
    @@{command judgment} @{syntax constdecl}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1118
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1119
    @@{attribute atomize} ('(' 'full' ')')?
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1120
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1121
    @@{attribute rule_format} ('(' 'noasm' ')')?
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1122
  "}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1123
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1124
  \begin{description}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1125
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1126
  \item @{command "judgment"}~@{text "c :: \<sigma> (mx)"} declares constant
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1127
  @{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
  1128
  type @{text \<sigma>} should specify a coercion of the category of
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1129
  object-level propositions to @{text prop} of the Pure meta-logic;
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1130
  the mixfix annotation @{text "(mx)"} would typically just link the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1131
  object language (internally of syntactic category @{text logic})
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1132
  with that of @{text prop}.  Only one @{command "judgment"}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1133
  declaration may be given in any theory development.
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1134
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1135
  \item @{method atomize} (as a method) rewrites any non-atomic
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1136
  premises of a sub-goal, using the meta-level equations declared via
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1137
  @{attribute atomize} (as an attribute) beforehand.  As a result,
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1138
  heavily nested goals become amenable to fundamental operations such
42626
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 42617
diff changeset
  1139
  as resolution (cf.\ the @{method (Pure) rule} method).  Giving the ``@{text
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1140
  "(full)"}'' option here means to turn the whole subgoal into an
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1141
  object-statement (if possible), including the outermost parameters
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1142
  and assumptions as well.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1143
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1144
  A typical collection of @{attribute atomize} rules for a particular
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1145
  object-logic would provide an internalization for each of the
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1146
  connectives of @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"}.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1147
  Meta-level conjunction should be covered as well (this is
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1148
  particularly important for locales, see \secref{sec:locale}).
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1149
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1150
  \item @{attribute rule_format} rewrites a theorem by the equalities
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1151
  declared as @{attribute rulify} rules in the current object-logic.
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1152
  By default, the result is fully normalized, including assumptions
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1153
  and conclusions at any depth.  The @{text "(no_asm)"} option
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1154
  restricts the transformation to the conclusion of a rule.
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1155
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1156
  In common object-logics (HOL, FOL, ZF), the effect of @{attribute
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1157
  rule_format} is to replace (bounded) universal quantification
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1158
  (@{text "\<forall>"}) and implication (@{text "\<longrightarrow>"}) by the corresponding
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1159
  rule statements over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1160
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1161
  \end{description}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1162
*}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1163
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1164
end