doc-src/IsarRef/Thy/Generic.thy
author wenzelm
Tue, 09 Aug 2011 15:41:00 +0200
changeset 44965 f7bbfdf4b4a7
parent 44238 3f1469de9e47
child 45782 884d27ede438
permissions -rw-r--r--
updated documentation of method "split" according to e6a4bb832b46;
wenzelm@26782
     1
theory Generic
wenzelm@43522
     2
imports Base Main
wenzelm@26782
     3
begin
wenzelm@26782
     4
wenzelm@26782
     5
chapter {* Generic tools and packages \label{ch:gen-tools} *}
wenzelm@26782
     6
wenzelm@43526
     7
section {* Configuration options \label{sec:config} *}
wenzelm@26782
     8
wenzelm@40537
     9
text {* Isabelle/Pure maintains a record of named configuration
wenzelm@40537
    10
  options within the theory or proof context, with values of type
wenzelm@40537
    11
  @{ML_type bool}, @{ML_type int}, @{ML_type real}, or @{ML_type
wenzelm@40537
    12
  string}.  Tools may declare options in ML, and then refer to these
wenzelm@40537
    13
  values (relative to the context).  Thus global reference variables
wenzelm@40537
    14
  are easily avoided.  The user may change the value of a
wenzelm@40537
    15
  configuration option by means of an associated attribute of the same
wenzelm@40537
    16
  name.  This form of context declaration works particularly well with
wenzelm@43526
    17
  commands such as @{command "declare"} or @{command "using"} like
wenzelm@43526
    18
  this:
wenzelm@43526
    19
*}
wenzelm@26782
    20
wenzelm@43526
    21
declare [[show_main_goal = false]]
wenzelm@43526
    22
wenzelm@43526
    23
notepad
wenzelm@43526
    24
begin
wenzelm@43526
    25
  note [[show_main_goal = true]]
wenzelm@43526
    26
end
wenzelm@43526
    27
wenzelm@43526
    28
text {* For historical reasons, some tools cannot take the full proof
wenzelm@26782
    29
  context into account and merely refer to the background theory.
wenzelm@26782
    30
  This is accommodated by configuration options being declared as
wenzelm@26782
    31
  ``global'', which may not be changed within a local context.
wenzelm@26782
    32
wenzelm@26782
    33
  \begin{matharray}{rcll}
wenzelm@28761
    34
    @{command_def "print_configs"} & : & @{text "context \<rightarrow>"} \\
wenzelm@26782
    35
  \end{matharray}
wenzelm@26782
    36
wenzelm@43467
    37
  @{rail "
wenzelm@43467
    38
    @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))?
wenzelm@43467
    39
  "}
wenzelm@26782
    40
wenzelm@28760
    41
  \begin{description}
wenzelm@26782
    42
  
wenzelm@28760
    43
  \item @{command "print_configs"} prints the available configuration
wenzelm@28760
    44
  options, with names, types, and current values.
wenzelm@26782
    45
  
wenzelm@28760
    46
  \item @{text "name = value"} as an attribute expression modifies the
wenzelm@28760
    47
  named option, with the syntax of the value depending on the option's
wenzelm@28760
    48
  type.  For @{ML_type bool} the default value is @{text true}.  Any
wenzelm@28760
    49
  attempt to change a global option in a local context is ignored.
wenzelm@26782
    50
wenzelm@28760
    51
  \end{description}
wenzelm@26782
    52
*}
wenzelm@26782
    53
wenzelm@26782
    54
wenzelm@27040
    55
section {* Basic proof tools *}
wenzelm@26782
    56
wenzelm@26782
    57
subsection {* Miscellaneous methods and attributes \label{sec:misc-meth-att} *}
wenzelm@26782
    58
wenzelm@26782
    59
text {*
wenzelm@26782
    60
  \begin{matharray}{rcl}
wenzelm@28761
    61
    @{method_def unfold} & : & @{text method} \\
wenzelm@28761
    62
    @{method_def fold} & : & @{text method} \\
wenzelm@28761
    63
    @{method_def insert} & : & @{text method} \\[0.5ex]
wenzelm@28761
    64
    @{method_def erule}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
    65
    @{method_def drule}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
    66
    @{method_def frule}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@44236
    67
    @{method_def intro} & : & @{text method} \\
wenzelm@44236
    68
    @{method_def elim} & : & @{text method} \\
wenzelm@28761
    69
    @{method_def succeed} & : & @{text method} \\
wenzelm@28761
    70
    @{method_def fail} & : & @{text method} \\
wenzelm@26782
    71
  \end{matharray}
wenzelm@26782
    72
wenzelm@43467
    73
  @{rail "
wenzelm@43467
    74
    (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thmrefs}
wenzelm@26782
    75
    ;
wenzelm@43467
    76
    (@@{method erule} | @@{method drule} | @@{method frule})
wenzelm@43467
    77
      ('(' @{syntax nat} ')')? @{syntax thmrefs}
wenzelm@44236
    78
    ;
wenzelm@44236
    79
    (@@{method intro} | @@{method elim}) @{syntax thmrefs}?
wenzelm@43467
    80
  "}
wenzelm@26782
    81
wenzelm@28760
    82
  \begin{description}
wenzelm@26782
    83
  
wenzelm@28760
    84
  \item @{method unfold}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} and @{method fold}~@{text
wenzelm@28760
    85
  "a\<^sub>1 \<dots> a\<^sub>n"} expand (or fold back) the given definitions throughout
wenzelm@28760
    86
  all goals; any chained facts provided are inserted into the goal and
wenzelm@28760
    87
  subject to rewriting as well.
wenzelm@26782
    88
wenzelm@28760
    89
  \item @{method insert}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} inserts theorems as facts
wenzelm@28760
    90
  into all goals of the proof state.  Note that current facts
wenzelm@28760
    91
  indicated for forward chaining are ignored.
wenzelm@26782
    92
wenzelm@30397
    93
  \item @{method erule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, @{method
wenzelm@30397
    94
  drule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, and @{method frule}~@{text
wenzelm@30397
    95
  "a\<^sub>1 \<dots> a\<^sub>n"} are similar to the basic @{method rule}
wenzelm@30397
    96
  method (see \secref{sec:pure-meth-att}), but apply rules by
wenzelm@30397
    97
  elim-resolution, destruct-resolution, and forward-resolution,
wenzelm@30397
    98
  respectively \cite{isabelle-implementation}.  The optional natural
wenzelm@30397
    99
  number argument (default 0) specifies additional assumption steps to
wenzelm@30397
   100
  be performed here.
wenzelm@26782
   101
wenzelm@26782
   102
  Note that these methods are improper ones, mainly serving for
wenzelm@26782
   103
  experimentation and tactic script emulation.  Different modes of
wenzelm@26782
   104
  basic rule application are usually expressed in Isar at the proof
wenzelm@26782
   105
  language level, rather than via implicit proof state manipulations.
wenzelm@26782
   106
  For example, a proper single-step elimination would be done using
wenzelm@26782
   107
  the plain @{method rule} method, with forward chaining of current
wenzelm@26782
   108
  facts.
wenzelm@26782
   109
wenzelm@44236
   110
  \item @{method intro} and @{method elim} repeatedly refine some goal
wenzelm@44236
   111
  by intro- or elim-resolution, after having inserted any chained
wenzelm@44236
   112
  facts.  Exactly the rules given as arguments are taken into account;
wenzelm@44236
   113
  this allows fine-tuned decomposition of a proof problem, in contrast
wenzelm@44236
   114
  to common automated tools.
wenzelm@44236
   115
wenzelm@28760
   116
  \item @{method succeed} yields a single (unchanged) result; it is
wenzelm@26782
   117
  the identity of the ``@{text ","}'' method combinator (cf.\
wenzelm@28754
   118
  \secref{sec:proof-meth}).
wenzelm@26782
   119
wenzelm@28760
   120
  \item @{method fail} yields an empty result sequence; it is the
wenzelm@26782
   121
  identity of the ``@{text "|"}'' method combinator (cf.\
wenzelm@28754
   122
  \secref{sec:proof-meth}).
wenzelm@26782
   123
wenzelm@28760
   124
  \end{description}
wenzelm@26782
   125
wenzelm@26782
   126
  \begin{matharray}{rcl}
wenzelm@28761
   127
    @{attribute_def tagged} & : & @{text attribute} \\
wenzelm@28761
   128
    @{attribute_def untagged} & : & @{text attribute} \\[0.5ex]
wenzelm@28761
   129
    @{attribute_def THEN} & : & @{text attribute} \\
wenzelm@28761
   130
    @{attribute_def COMP} & : & @{text attribute} \\[0.5ex]
wenzelm@28761
   131
    @{attribute_def unfolded} & : & @{text attribute} \\
wenzelm@28761
   132
    @{attribute_def folded} & : & @{text attribute} \\[0.5ex]
wenzelm@28761
   133
    @{attribute_def rotated} & : & @{text attribute} \\
wenzelm@28761
   134
    @{attribute_def (Pure) elim_format} & : & @{text attribute} \\
wenzelm@28761
   135
    @{attribute_def standard}@{text "\<^sup>*"} & : & @{text attribute} \\
wenzelm@28761
   136
    @{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\
wenzelm@26782
   137
  \end{matharray}
wenzelm@26782
   138
wenzelm@43467
   139
  @{rail "
wenzelm@43467
   140
    @@{attribute tagged} @{syntax name} @{syntax name}
wenzelm@26782
   141
    ;
wenzelm@43467
   142
    @@{attribute untagged} @{syntax name}
wenzelm@26782
   143
    ;
wenzelm@43467
   144
    (@@{attribute THEN} | @@{attribute COMP}) ('[' @{syntax nat} ']')? @{syntax thmref}
wenzelm@26782
   145
    ;
wenzelm@43467
   146
    (@@{attribute unfolded} | @@{attribute folded}) @{syntax thmrefs}
wenzelm@26782
   147
    ;
wenzelm@43467
   148
    @@{attribute rotated} @{syntax int}?
wenzelm@43467
   149
  "}
wenzelm@26782
   150
wenzelm@28760
   151
  \begin{description}
wenzelm@26782
   152
wenzelm@28764
   153
  \item @{attribute tagged}~@{text "name value"} and @{attribute
wenzelm@28760
   154
  untagged}~@{text name} add and remove \emph{tags} of some theorem.
wenzelm@26782
   155
  Tags may be any list of string pairs that serve as formal comment.
wenzelm@28764
   156
  The first string is considered the tag name, the second its value.
wenzelm@28764
   157
  Note that @{attribute untagged} removes any tags of the same name.
wenzelm@26782
   158
wenzelm@28760
   159
  \item @{attribute THEN}~@{text a} and @{attribute COMP}~@{text a}
wenzelm@26782
   160
  compose rules by resolution.  @{attribute THEN} resolves with the
wenzelm@26782
   161
  first premise of @{text a} (an alternative position may be also
wenzelm@26782
   162
  specified); the @{attribute COMP} version skips the automatic
wenzelm@30462
   163
  lifting process that is normally intended (cf.\ @{ML "op RS"} and
wenzelm@30462
   164
  @{ML "op COMP"} in \cite{isabelle-implementation}).
wenzelm@26782
   165
  
wenzelm@28760
   166
  \item @{attribute unfolded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} and @{attribute
wenzelm@28760
   167
  folded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} expand and fold back again the given
wenzelm@28760
   168
  definitions throughout a rule.
wenzelm@26782
   169
wenzelm@28760
   170
  \item @{attribute rotated}~@{text n} rotate the premises of a
wenzelm@26782
   171
  theorem by @{text n} (default 1).
wenzelm@26782
   172
wenzelm@28760
   173
  \item @{attribute (Pure) elim_format} turns a destruction rule into
wenzelm@28760
   174
  elimination rule format, by resolving with the rule @{prop "PROP A \<Longrightarrow>
wenzelm@28760
   175
  (PROP A \<Longrightarrow> PROP B) \<Longrightarrow> PROP B"}.
wenzelm@26782
   176
  
wenzelm@26782
   177
  Note that the Classical Reasoner (\secref{sec:classical}) provides
wenzelm@26782
   178
  its own version of this operation.
wenzelm@26782
   179
wenzelm@28760
   180
  \item @{attribute standard} puts a theorem into the standard form of
wenzelm@28760
   181
  object-rules at the outermost theory level.  Note that this
wenzelm@26782
   182
  operation violates the local proof context (including active
wenzelm@26782
   183
  locales).
wenzelm@26782
   184
wenzelm@28760
   185
  \item @{attribute no_vars} replaces schematic variables by free
wenzelm@26782
   186
  ones; this is mainly for tuning output of pretty printed theorems.
wenzelm@26782
   187
wenzelm@28760
   188
  \end{description}
wenzelm@26782
   189
*}
wenzelm@26782
   190
wenzelm@26782
   191
wenzelm@27044
   192
subsection {* Low-level equational reasoning *}
wenzelm@27044
   193
wenzelm@27044
   194
text {*
wenzelm@27044
   195
  \begin{matharray}{rcl}
wenzelm@28761
   196
    @{method_def subst} & : & @{text method} \\
wenzelm@28761
   197
    @{method_def hypsubst} & : & @{text method} \\
wenzelm@28761
   198
    @{method_def split} & : & @{text method} \\
wenzelm@27044
   199
  \end{matharray}
wenzelm@27044
   200
wenzelm@43467
   201
  @{rail "
wenzelm@43575
   202
    @@{method subst} ('(' 'asm' ')')? \\ ('(' (@{syntax nat}+) ')')? @{syntax thmref}
wenzelm@27044
   203
    ;
wenzelm@44965
   204
    @@{method split} @{syntax thmrefs}
wenzelm@43467
   205
  "}
wenzelm@27044
   206
wenzelm@27044
   207
  These methods provide low-level facilities for equational reasoning
wenzelm@27044
   208
  that are intended for specialized applications only.  Normally,
wenzelm@27044
   209
  single step calculations would be performed in a structured text
wenzelm@27044
   210
  (see also \secref{sec:calculation}), while the Simplifier methods
wenzelm@27044
   211
  provide the canonical way for automated normalization (see
wenzelm@27044
   212
  \secref{sec:simplifier}).
wenzelm@27044
   213
wenzelm@28760
   214
  \begin{description}
wenzelm@27044
   215
wenzelm@28760
   216
  \item @{method subst}~@{text eq} performs a single substitution step
wenzelm@28760
   217
  using rule @{text eq}, which may be either a meta or object
wenzelm@27044
   218
  equality.
wenzelm@27044
   219
wenzelm@28760
   220
  \item @{method subst}~@{text "(asm) eq"} substitutes in an
wenzelm@27044
   221
  assumption.
wenzelm@27044
   222
wenzelm@28760
   223
  \item @{method subst}~@{text "(i \<dots> j) eq"} performs several
wenzelm@27044
   224
  substitutions in the conclusion. The numbers @{text i} to @{text j}
wenzelm@27044
   225
  indicate the positions to substitute at.  Positions are ordered from
wenzelm@27044
   226
  the top of the term tree moving down from left to right. For
wenzelm@27044
   227
  example, in @{text "(a + b) + (c + d)"} there are three positions
wenzelm@28760
   228
  where commutativity of @{text "+"} is applicable: 1 refers to @{text
wenzelm@28760
   229
  "a + b"}, 2 to the whole term, and 3 to @{text "c + d"}.
wenzelm@27044
   230
wenzelm@27044
   231
  If the positions in the list @{text "(i \<dots> j)"} are non-overlapping
wenzelm@27044
   232
  (e.g.\ @{text "(2 3)"} in @{text "(a + b) + (c + d)"}) you may
wenzelm@27044
   233
  assume all substitutions are performed simultaneously.  Otherwise
wenzelm@27044
   234
  the behaviour of @{text subst} is not specified.
wenzelm@27044
   235
wenzelm@28760
   236
  \item @{method subst}~@{text "(asm) (i \<dots> j) eq"} performs the
wenzelm@27071
   237
  substitutions in the assumptions. The positions refer to the
wenzelm@27071
   238
  assumptions in order from left to right.  For example, given in a
wenzelm@27071
   239
  goal of the form @{text "P (a + b) \<Longrightarrow> P (c + d) \<Longrightarrow> \<dots>"}, position 1 of
wenzelm@27071
   240
  commutativity of @{text "+"} is the subterm @{text "a + b"} and
wenzelm@27071
   241
  position 2 is the subterm @{text "c + d"}.
wenzelm@27044
   242
wenzelm@28760
   243
  \item @{method hypsubst} performs substitution using some
wenzelm@27044
   244
  assumption; this only works for equations of the form @{text "x =
wenzelm@27044
   245
  t"} where @{text x} is a free or bound variable.
wenzelm@27044
   246
wenzelm@28760
   247
  \item @{method split}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} performs single-step case
wenzelm@44965
   248
  splitting using the given rules.  Splitting is performed in the
wenzelm@44965
   249
  conclusion or some assumption of the subgoal, depending of the
wenzelm@44965
   250
  structure of the rule.
wenzelm@27044
   251
  
wenzelm@27044
   252
  Note that the @{method simp} method already involves repeated
wenzelm@44965
   253
  application of split rules as declared in the current context, using
wenzelm@44965
   254
  @{attribute split}, for example.
wenzelm@27044
   255
wenzelm@28760
   256
  \end{description}
wenzelm@27044
   257
*}
wenzelm@27044
   258
wenzelm@27044
   259
wenzelm@26782
   260
subsection {* Further tactic emulations \label{sec:tactics} *}
wenzelm@26782
   261
wenzelm@26782
   262
text {*
wenzelm@26782
   263
  The following improper proof methods emulate traditional tactics.
wenzelm@26782
   264
  These admit direct access to the goal state, which is normally
wenzelm@26782
   265
  considered harmful!  In particular, this may involve both numbered
wenzelm@26782
   266
  goal addressing (default 1), and dynamic instantiation within the
wenzelm@26782
   267
  scope of some subgoal.
wenzelm@26782
   268
wenzelm@26782
   269
  \begin{warn}
wenzelm@26782
   270
    Dynamic instantiations refer to universally quantified parameters
wenzelm@26782
   271
    of a subgoal (the dynamic context) rather than fixed variables and
wenzelm@26782
   272
    term abbreviations of a (static) Isar context.
wenzelm@26782
   273
  \end{warn}
wenzelm@26782
   274
wenzelm@26782
   275
  Tactic emulation methods, unlike their ML counterparts, admit
wenzelm@26782
   276
  simultaneous instantiation from both dynamic and static contexts.
wenzelm@26782
   277
  If names occur in both contexts goal parameters hide locally fixed
wenzelm@26782
   278
  variables.  Likewise, schematic variables refer to term
wenzelm@26782
   279
  abbreviations, if present in the static context.  Otherwise the
wenzelm@26782
   280
  schematic variable is interpreted as a schematic variable and left
wenzelm@26782
   281
  to be solved by unification with certain parts of the subgoal.
wenzelm@26782
   282
wenzelm@26782
   283
  Note that the tactic emulation proof methods in Isabelle/Isar are
wenzelm@26782
   284
  consistently named @{text foo_tac}.  Note also that variable names
wenzelm@26782
   285
  occurring on left hand sides of instantiations must be preceded by a
wenzelm@26782
   286
  question mark if they coincide with a keyword or contain dots.  This
wenzelm@26782
   287
  is consistent with the attribute @{attribute "where"} (see
wenzelm@26782
   288
  \secref{sec:pure-meth-att}).
wenzelm@26782
   289
wenzelm@26782
   290
  \begin{matharray}{rcl}
wenzelm@28761
   291
    @{method_def rule_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   292
    @{method_def erule_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   293
    @{method_def drule_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   294
    @{method_def frule_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   295
    @{method_def cut_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   296
    @{method_def thin_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   297
    @{method_def subgoal_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   298
    @{method_def rename_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   299
    @{method_def rotate_tac}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   300
    @{method_def tactic}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@28761
   301
    @{method_def raw_tactic}@{text "\<^sup>*"} & : & @{text method} \\
wenzelm@26782
   302
  \end{matharray}
wenzelm@26782
   303
wenzelm@43467
   304
  @{rail "
wenzelm@43467
   305
    (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} |
wenzelm@43576
   306
      @@{method frule_tac} | @@{method cut_tac} | @@{method thin_tac}) @{syntax goal_spec}? \\
wenzelm@43488
   307
    ( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} )
wenzelm@26782
   308
    ;
wenzelm@43576
   309
    @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +)
wenzelm@26782
   310
    ;
wenzelm@43576
   311
    @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +)
wenzelm@26782
   312
    ;
wenzelm@43576
   313
    @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}?
wenzelm@26782
   314
    ;
wenzelm@43467
   315
    (@@{method tactic} | @@{method raw_tactic}) @{syntax text}
wenzelm@26782
   316
    ;
wenzelm@26782
   317
wenzelm@43488
   318
    dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and')
wenzelm@43488
   319
  "}
wenzelm@26782
   320
wenzelm@28760
   321
\begin{description}
wenzelm@26782
   322
wenzelm@28760
   323
  \item @{method rule_tac} etc. do resolution of rules with explicit
wenzelm@26782
   324
  instantiation.  This works the same way as the ML tactics @{ML
wenzelm@30397
   325
  res_inst_tac} etc. (see \cite{isabelle-implementation})
wenzelm@26782
   326
wenzelm@26782
   327
  Multiple rules may be only given if there is no instantiation; then
wenzelm@26782
   328
  @{method rule_tac} is the same as @{ML resolve_tac} in ML (see
wenzelm@30397
   329
  \cite{isabelle-implementation}).
wenzelm@26782
   330
wenzelm@28760
   331
  \item @{method cut_tac} inserts facts into the proof state as
wenzelm@27209
   332
  assumption of a subgoal, see also @{ML Tactic.cut_facts_tac} in
wenzelm@30397
   333
  \cite{isabelle-implementation}.  Note that the scope of schematic
wenzelm@26782
   334
  variables is spread over the main goal statement.  Instantiations
wenzelm@28760
   335
  may be given as well, see also ML tactic @{ML cut_inst_tac} in
wenzelm@30397
   336
  \cite{isabelle-implementation}.
wenzelm@26782
   337
wenzelm@28760
   338
  \item @{method thin_tac}~@{text \<phi>} deletes the specified assumption
wenzelm@28760
   339
  from a subgoal; note that @{text \<phi>} may contain schematic variables.
wenzelm@30397
   340
  See also @{ML thin_tac} in \cite{isabelle-implementation}.
wenzelm@28760
   341
wenzelm@28760
   342
  \item @{method subgoal_tac}~@{text \<phi>} adds @{text \<phi>} as an
wenzelm@27239
   343
  assumption to a subgoal.  See also @{ML subgoal_tac} and @{ML
wenzelm@30397
   344
  subgoals_tac} in \cite{isabelle-implementation}.
wenzelm@26782
   345
wenzelm@28760
   346
  \item @{method rename_tac}~@{text "x\<^sub>1 \<dots> x\<^sub>n"} renames parameters of a
wenzelm@28760
   347
  goal according to the list @{text "x\<^sub>1, \<dots>, x\<^sub>n"}, which refers to the
wenzelm@28760
   348
  \emph{suffix} of variables.
wenzelm@26782
   349
wenzelm@28760
   350
  \item @{method rotate_tac}~@{text n} rotates the assumptions of a
wenzelm@26782
   351
  goal by @{text n} positions: from right to left if @{text n} is
wenzelm@26782
   352
  positive, and from left to right if @{text n} is negative; the
wenzelm@26782
   353
  default value is 1.  See also @{ML rotate_tac} in
wenzelm@30397
   354
  \cite{isabelle-implementation}.
wenzelm@26782
   355
wenzelm@28760
   356
  \item @{method tactic}~@{text "text"} produces a proof method from
wenzelm@26782
   357
  any ML text of type @{ML_type tactic}.  Apart from the usual ML
wenzelm@27223
   358
  environment and the current proof context, the ML code may refer to
wenzelm@27223
   359
  the locally bound values @{ML_text facts}, which indicates any
wenzelm@27223
   360
  current facts used for forward-chaining.
wenzelm@26782
   361
wenzelm@28760
   362
  \item @{method raw_tactic} is similar to @{method tactic}, but
wenzelm@27223
   363
  presents the goal state in its raw internal form, where simultaneous
wenzelm@27223
   364
  subgoals appear as conjunction of the logical framework instead of
wenzelm@27223
   365
  the usual split into several subgoals.  While feature this is useful
wenzelm@27223
   366
  for debugging of complex method definitions, it should not never
wenzelm@27223
   367
  appear in production theories.
wenzelm@26782
   368
wenzelm@28760
   369
  \end{description}
wenzelm@26782
   370
*}
wenzelm@26782
   371
wenzelm@26782
   372
wenzelm@27040
   373
section {* The Simplifier \label{sec:simplifier} *}
wenzelm@26782
   374
wenzelm@27040
   375
subsection {* Simplification methods *}
wenzelm@26782
   376
wenzelm@26782
   377
text {*
wenzelm@26782
   378
  \begin{matharray}{rcl}
wenzelm@28761
   379
    @{method_def simp} & : & @{text method} \\
wenzelm@28761
   380
    @{method_def simp_all} & : & @{text method} \\
wenzelm@26782
   381
  \end{matharray}
wenzelm@26782
   382
wenzelm@43467
   383
  @{rail "
wenzelm@43467
   384
    (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * )
wenzelm@26782
   385
    ;
wenzelm@26782
   386
wenzelm@40516
   387
    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')'
wenzelm@26782
   388
    ;
wenzelm@43467
   389
    @{syntax_def simpmod}: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') |
wenzelm@43467
   390
      'split' (() | 'add' | 'del')) ':' @{syntax thmrefs}
wenzelm@43467
   391
  "}
wenzelm@26782
   392
wenzelm@28760
   393
  \begin{description}
wenzelm@26782
   394
wenzelm@28760
   395
  \item @{method simp} invokes the Simplifier, after declaring
wenzelm@26782
   396
  additional rules according to the arguments given.  Note that the
wenzelm@43467
   397
  @{text only} modifier first removes all other rewrite rules,
wenzelm@26782
   398
  congruences, and looper tactics (including splits), and then behaves
wenzelm@43467
   399
  like @{text add}.
wenzelm@26782
   400
wenzelm@43467
   401
  \medskip The @{text cong} modifiers add or delete Simplifier
wenzelm@26782
   402
  congruence rules (see also \cite{isabelle-ref}), the default is to
wenzelm@26782
   403
  add.
wenzelm@26782
   404
wenzelm@43467
   405
  \medskip The @{text split} modifiers add or delete rules for the
wenzelm@26782
   406
  Splitter (see also \cite{isabelle-ref}), the default is to add.
wenzelm@26782
   407
  This works only if the Simplifier method has been properly setup to
wenzelm@26782
   408
  include the Splitter (all major object logics such HOL, HOLCF, FOL,
wenzelm@26782
   409
  ZF do this already).
wenzelm@26782
   410
wenzelm@28760
   411
  \item @{method simp_all} is similar to @{method simp}, but acts on
wenzelm@26782
   412
  all goals (backwards from the last to the first one).
wenzelm@26782
   413
wenzelm@28760
   414
  \end{description}
wenzelm@26782
   415
wenzelm@26782
   416
  By default the Simplifier methods take local assumptions fully into
wenzelm@26782
   417
  account, using equational assumptions in the subsequent
wenzelm@26782
   418
  normalization process, or simplifying assumptions themselves (cf.\
wenzelm@30397
   419
  @{ML asm_full_simp_tac} in \cite{isabelle-ref}).  In structured
wenzelm@30397
   420
  proofs this is usually quite well behaved in practice: just the
wenzelm@30397
   421
  local premises of the actual goal are involved, additional facts may
wenzelm@30397
   422
  be inserted via explicit forward-chaining (via @{command "then"},
wenzelm@35613
   423
  @{command "from"}, @{command "using"} etc.).
wenzelm@26782
   424
wenzelm@26782
   425
  Additional Simplifier options may be specified to tune the behavior
wenzelm@26782
   426
  further (mostly for unstructured scripts with many accidental local
wenzelm@26782
   427
  facts): ``@{text "(no_asm)"}'' means assumptions are ignored
wenzelm@26782
   428
  completely (cf.\ @{ML simp_tac}), ``@{text "(no_asm_simp)"}'' means
wenzelm@26782
   429
  assumptions are used in the simplification of the conclusion but are
wenzelm@26782
   430
  not themselves simplified (cf.\ @{ML asm_simp_tac}), and ``@{text
wenzelm@26782
   431
  "(no_asm_use)"}'' means assumptions are simplified but are not used
wenzelm@26782
   432
  in the simplification of each other or the conclusion (cf.\ @{ML
wenzelm@26782
   433
  full_simp_tac}).  For compatibility reasons, there is also an option
wenzelm@26782
   434
  ``@{text "(asm_lr)"}'', which means that an assumption is only used
wenzelm@26782
   435
  for simplifying assumptions which are to the right of it (cf.\ @{ML
wenzelm@26782
   436
  asm_lr_simp_tac}).
wenzelm@26782
   437
wenzelm@27092
   438
  The configuration option @{text "depth_limit"} limits the number of
wenzelm@26782
   439
  recursive invocations of the simplifier during conditional
wenzelm@26782
   440
  rewriting.
wenzelm@26782
   441
wenzelm@26782
   442
  \medskip The Splitter package is usually configured to work as part
wenzelm@26782
   443
  of the Simplifier.  The effect of repeatedly applying @{ML
wenzelm@26782
   444
  split_tac} can be simulated by ``@{text "(simp only: split:
wenzelm@26782
   445
  a\<^sub>1 \<dots> a\<^sub>n)"}''.  There is also a separate @{text split}
wenzelm@26782
   446
  method available for single-step case splitting.
wenzelm@26782
   447
*}
wenzelm@26782
   448
wenzelm@26782
   449
wenzelm@27040
   450
subsection {* Declaring rules *}
wenzelm@26782
   451
wenzelm@26782
   452
text {*
wenzelm@26782
   453
  \begin{matharray}{rcl}
wenzelm@28761
   454
    @{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
wenzelm@28761
   455
    @{attribute_def simp} & : & @{text attribute} \\
wenzelm@28761
   456
    @{attribute_def cong} & : & @{text attribute} \\
wenzelm@28761
   457
    @{attribute_def split} & : & @{text attribute} \\
wenzelm@26782
   458
  \end{matharray}
wenzelm@26782
   459
wenzelm@43467
   460
  @{rail "
wenzelm@43467
   461
    (@@{attribute simp} | @@{attribute cong} | @@{attribute split}) (() | 'add' | 'del')
wenzelm@43467
   462
  "}
wenzelm@26782
   463
wenzelm@28760
   464
  \begin{description}
wenzelm@26782
   465
wenzelm@28760
   466
  \item @{command "print_simpset"} prints the collection of rules
wenzelm@26782
   467
  declared to the Simplifier, which is also known as ``simpset''
wenzelm@26782
   468
  internally \cite{isabelle-ref}.
wenzelm@26782
   469
wenzelm@28760
   470
  \item @{attribute simp} declares simplification rules.
wenzelm@26782
   471
wenzelm@28760
   472
  \item @{attribute cong} declares congruence rules.
wenzelm@26782
   473
wenzelm@28760
   474
  \item @{attribute split} declares case split rules.
wenzelm@26782
   475
wenzelm@28760
   476
  \end{description}
wenzelm@26782
   477
*}
wenzelm@26782
   478
wenzelm@26782
   479
wenzelm@27040
   480
subsection {* Simplification procedures *}
wenzelm@26782
   481
wenzelm@44129
   482
text {* Simplification procedures are ML functions that produce proven
wenzelm@44129
   483
  rewrite rules on demand.  They are associated with higher-order
wenzelm@44129
   484
  patterns that approximate the left-hand sides of equations.  The
wenzelm@44129
   485
  Simplifier first matches the current redex against one of the LHS
wenzelm@44129
   486
  patterns; if this succeeds, the corresponding ML function is
wenzelm@44129
   487
  invoked, passing the Simplifier context and redex term.  Thus rules
wenzelm@44129
   488
  may be specifically fashioned for particular situations, resulting
wenzelm@44129
   489
  in a more powerful mechanism than term rewriting by a fixed set of
wenzelm@44129
   490
  rules.
wenzelm@44129
   491
wenzelm@44129
   492
  Any successful result needs to be a (possibly conditional) rewrite
wenzelm@44129
   493
  rule @{text "t \<equiv> u"} that is applicable to the current redex.  The
wenzelm@44129
   494
  rule will be applied just as any ordinary rewrite rule.  It is
wenzelm@44129
   495
  expected to be already in \emph{internal form}, bypassing the
wenzelm@44129
   496
  automatic preprocessing of object-level equivalences.
wenzelm@44129
   497
wenzelm@26782
   498
  \begin{matharray}{rcl}
wenzelm@28761
   499
    @{command_def "simproc_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
wenzelm@28761
   500
    simproc & : & @{text attribute} \\
wenzelm@26782
   501
  \end{matharray}
wenzelm@26782
   502
wenzelm@43467
   503
  @{rail "
wenzelm@43467
   504
    @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '='
wenzelm@43467
   505
      @{syntax text} \\ (@'identifier' (@{syntax nameref}+))?
wenzelm@26782
   506
    ;
wenzelm@26782
   507
wenzelm@43467
   508
    @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+)
wenzelm@43467
   509
  "}
wenzelm@26782
   510
wenzelm@28760
   511
  \begin{description}
wenzelm@26782
   512
wenzelm@28760
   513
  \item @{command "simproc_setup"} defines a named simplification
wenzelm@26782
   514
  procedure that is invoked by the Simplifier whenever any of the
wenzelm@26782
   515
  given term patterns match the current redex.  The implementation,
wenzelm@26782
   516
  which is provided as ML source text, needs to be of type @{ML_type
wenzelm@26782
   517
  "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type
wenzelm@26782
   518
  cterm} represents the current redex @{text r} and the result is
wenzelm@26782
   519
  supposed to be some proven rewrite rule @{text "r \<equiv> r'"} (or a
wenzelm@26782
   520
  generalized version), or @{ML NONE} to indicate failure.  The
wenzelm@26782
   521
  @{ML_type simpset} argument holds the full context of the current
wenzelm@26782
   522
  Simplifier invocation, including the actual Isar proof context.  The
wenzelm@26782
   523
  @{ML_type morphism} informs about the difference of the original
wenzelm@26782
   524
  compilation context wrt.\ the one of the actual application later
wenzelm@26782
   525
  on.  The optional @{keyword "identifier"} specifies theorems that
wenzelm@26782
   526
  represent the logical content of the abstract theory of this
wenzelm@26782
   527
  simproc.
wenzelm@26782
   528
wenzelm@26782
   529
  Morphisms and identifiers are only relevant for simprocs that are
wenzelm@26782
   530
  defined within a local target context, e.g.\ in a locale.
wenzelm@26782
   531
wenzelm@28760
   532
  \item @{text "simproc add: name"} and @{text "simproc del: name"}
wenzelm@26782
   533
  add or delete named simprocs to the current Simplifier context.  The
wenzelm@26782
   534
  default is to add a simproc.  Note that @{command "simproc_setup"}
wenzelm@26782
   535
  already adds the new simproc to the subsequent context.
wenzelm@26782
   536
wenzelm@28760
   537
  \end{description}
wenzelm@26782
   538
*}
wenzelm@26782
   539
wenzelm@26782
   540
wenzelm@44129
   541
subsubsection {* Example *}
wenzelm@44129
   542
wenzelm@44129
   543
text {* The following simplification procedure for @{thm
wenzelm@44129
   544
  [source=false, show_types] unit_eq} in HOL performs fine-grained
wenzelm@44129
   545
  control over rule application, beyond higher-order pattern matching.
wenzelm@44129
   546
  Declaring @{thm unit_eq} as @{attribute simp} directly would make
wenzelm@44129
   547
  the simplifier loop!  Note that a version of this simplification
wenzelm@44129
   548
  procedure is already active in Isabelle/HOL.  *}
wenzelm@44129
   549
wenzelm@44129
   550
simproc_setup unit ("x::unit") = {*
wenzelm@44129
   551
  fn _ => fn _ => fn ct =>
wenzelm@44129
   552
    if HOLogic.is_unit (term_of ct) then NONE
wenzelm@44129
   553
    else SOME (mk_meta_eq @{thm unit_eq})
wenzelm@44129
   554
*}
wenzelm@44129
   555
wenzelm@44129
   556
text {* Since the Simplifier applies simplification procedures
wenzelm@44129
   557
  frequently, it is important to make the failure check in ML
wenzelm@44129
   558
  reasonably fast. *}
wenzelm@44129
   559
wenzelm@44129
   560
wenzelm@27040
   561
subsection {* Forward simplification *}
wenzelm@26782
   562
wenzelm@26782
   563
text {*
wenzelm@26782
   564
  \begin{matharray}{rcl}
wenzelm@28761
   565
    @{attribute_def simplified} & : & @{text attribute} \\
wenzelm@26782
   566
  \end{matharray}
wenzelm@26782
   567
wenzelm@43467
   568
  @{rail "
wenzelm@43467
   569
    @@{attribute simplified} opt? @{syntax thmrefs}?
wenzelm@26782
   570
    ;
wenzelm@26782
   571
wenzelm@40516
   572
    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')'
wenzelm@43467
   573
  "}
wenzelm@26782
   574
wenzelm@28760
   575
  \begin{description}
wenzelm@26782
   576
  
wenzelm@28760
   577
  \item @{attribute simplified}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} causes a theorem to
wenzelm@28760
   578
  be simplified, either by exactly the specified rules @{text "a\<^sub>1, \<dots>,
wenzelm@28760
   579
  a\<^sub>n"}, or the implicit Simplifier context if no arguments are given.
wenzelm@28760
   580
  The result is fully simplified by default, including assumptions and
wenzelm@28760
   581
  conclusion; the options @{text no_asm} etc.\ tune the Simplifier in
wenzelm@28760
   582
  the same way as the for the @{text simp} method.
wenzelm@26782
   583
wenzelm@26782
   584
  Note that forward simplification restricts the simplifier to its
wenzelm@26782
   585
  most basic operation of term rewriting; solver and looper tactics
wenzelm@26782
   586
  \cite{isabelle-ref} are \emph{not} involved here.  The @{text
wenzelm@26782
   587
  simplified} attribute should be only rarely required under normal
wenzelm@26782
   588
  circumstances.
wenzelm@26782
   589
wenzelm@28760
   590
  \end{description}
wenzelm@26782
   591
*}
wenzelm@26782
   592
wenzelm@26782
   593
wenzelm@27040
   594
section {* The Classical Reasoner \label{sec:classical} *}
wenzelm@26782
   595
wenzelm@44134
   596
subsection {* Basic concepts *}
wenzelm@44131
   597
wenzelm@44131
   598
text {* Although Isabelle is generic, many users will be working in
wenzelm@44131
   599
  some extension of classical first-order logic.  Isabelle/ZF is built
wenzelm@44131
   600
  upon theory FOL, while Isabelle/HOL conceptually contains
wenzelm@44131
   601
  first-order logic as a fragment.  Theorem-proving in predicate logic
wenzelm@44131
   602
  is undecidable, but many automated strategies have been developed to
wenzelm@44131
   603
  assist in this task.
wenzelm@44131
   604
wenzelm@44131
   605
  Isabelle's classical reasoner is a generic package that accepts
wenzelm@44131
   606
  certain information about a logic and delivers a suite of automatic
wenzelm@44131
   607
  proof tools, based on rules that are classified and declared in the
wenzelm@44131
   608
  context.  These proof procedures are slow and simplistic compared
wenzelm@44131
   609
  with high-end automated theorem provers, but they can save
wenzelm@44131
   610
  considerable time and effort in practice.  They can prove theorems
wenzelm@44131
   611
  such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few
wenzelm@44131
   612
  milliseconds (including full proof reconstruction): *}
wenzelm@44131
   613
wenzelm@44131
   614
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)"
wenzelm@44131
   615
  by blast
wenzelm@44131
   616
wenzelm@44131
   617
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)"
wenzelm@44131
   618
  by blast
wenzelm@44131
   619
wenzelm@44131
   620
text {* The proof tools are generic.  They are not restricted to
wenzelm@44131
   621
  first-order logic, and have been heavily used in the development of
wenzelm@44131
   622
  the Isabelle/HOL library and applications.  The tactics can be
wenzelm@44131
   623
  traced, and their components can be called directly; in this manner,
wenzelm@44131
   624
  any proof can be viewed interactively.  *}
wenzelm@44131
   625
wenzelm@44131
   626
wenzelm@44131
   627
subsubsection {* The sequent calculus *}
wenzelm@44131
   628
wenzelm@44131
   629
text {* Isabelle supports natural deduction, which is easy to use for
wenzelm@44131
   630
  interactive proof.  But natural deduction does not easily lend
wenzelm@44131
   631
  itself to automation, and has a bias towards intuitionism.  For
wenzelm@44131
   632
  certain proofs in classical logic, it can not be called natural.
wenzelm@44131
   633
  The \emph{sequent calculus}, a generalization of natural deduction,
wenzelm@44131
   634
  is easier to automate.
wenzelm@44131
   635
wenzelm@44131
   636
  A \textbf{sequent} has the form @{text "\<Gamma> \<turnstile> \<Delta>"}, where @{text "\<Gamma>"}
wenzelm@44131
   637
  and @{text "\<Delta>"} are sets of formulae.\footnote{For first-order
wenzelm@44131
   638
  logic, sequents can equivalently be made from lists or multisets of
wenzelm@44131
   639
  formulae.} The sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} is
wenzelm@44131
   640
  \textbf{valid} if @{text "P\<^sub>1 \<and> \<dots> \<and> P\<^sub>m"} implies @{text "Q\<^sub>1 \<or> \<dots> \<or>
wenzelm@44131
   641
  Q\<^sub>n"}.  Thus @{text "P\<^sub>1, \<dots>, P\<^sub>m"} represent assumptions, each of which
wenzelm@44131
   642
  is true, while @{text "Q\<^sub>1, \<dots>, Q\<^sub>n"} represent alternative goals.  A
wenzelm@44131
   643
  sequent is \textbf{basic} if its left and right sides have a common
wenzelm@44131
   644
  formula, as in @{text "P, Q \<turnstile> Q, R"}; basic sequents are trivially
wenzelm@44131
   645
  valid.
wenzelm@44131
   646
wenzelm@44131
   647
  Sequent rules are classified as \textbf{right} or \textbf{left},
wenzelm@44131
   648
  indicating which side of the @{text "\<turnstile>"} symbol they operate on.
wenzelm@44131
   649
  Rules that operate on the right side are analogous to natural
wenzelm@44131
   650
  deduction's introduction rules, and left rules are analogous to
wenzelm@44131
   651
  elimination rules.  The sequent calculus analogue of @{text "(\<longrightarrow>I)"}
wenzelm@44131
   652
  is the rule
wenzelm@44131
   653
  \[
wenzelm@44131
   654
  \infer[@{text "(\<longrightarrow>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<longrightarrow> Q"}}{@{text "P, \<Gamma> \<turnstile> \<Delta>, Q"}}
wenzelm@44131
   655
  \]
wenzelm@44131
   656
  Applying the rule backwards, this breaks down some implication on
wenzelm@44131
   657
  the right side of a sequent; @{text "\<Gamma>"} and @{text "\<Delta>"} stand for
wenzelm@44131
   658
  the sets of formulae that are unaffected by the inference.  The
wenzelm@44131
   659
  analogue of the pair @{text "(\<or>I1)"} and @{text "(\<or>I2)"} is the
wenzelm@44131
   660
  single rule
wenzelm@44131
   661
  \[
wenzelm@44131
   662
  \infer[@{text "(\<or>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<or> Q"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P, Q"}}
wenzelm@44131
   663
  \]
wenzelm@44131
   664
  This breaks down some disjunction on the right side, replacing it by
wenzelm@44131
   665
  both disjuncts.  Thus, the sequent calculus is a kind of
wenzelm@44131
   666
  multiple-conclusion logic.
wenzelm@44131
   667
wenzelm@44131
   668
  To illustrate the use of multiple formulae on the right, let us
wenzelm@44131
   669
  prove the classical theorem @{text "(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}.  Working
wenzelm@44131
   670
  backwards, we reduce this formula to a basic sequent:
wenzelm@44131
   671
  \[
wenzelm@44131
   672
  \infer[@{text "(\<or>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}}
wenzelm@44131
   673
    {\infer[@{text "(\<longrightarrow>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q), (Q \<longrightarrow> P)"}}
wenzelm@44131
   674
      {\infer[@{text "(\<longrightarrow>R)"}]{@{text "P \<turnstile> Q, (Q \<longrightarrow> P)"}}
wenzelm@44131
   675
        {@{text "P, Q \<turnstile> Q, P"}}}}
wenzelm@44131
   676
  \]
wenzelm@44131
   677
wenzelm@44131
   678
  This example is typical of the sequent calculus: start with the
wenzelm@44131
   679
  desired theorem and apply rules backwards in a fairly arbitrary
wenzelm@44131
   680
  manner.  This yields a surprisingly effective proof procedure.
wenzelm@44131
   681
  Quantifiers add only few complications, since Isabelle handles
wenzelm@44131
   682
  parameters and schematic variables.  See \cite[Chapter
wenzelm@44131
   683
  10]{paulson-ml2} for further discussion.  *}
wenzelm@44131
   684
wenzelm@44131
   685
wenzelm@44131
   686
subsubsection {* Simulating sequents by natural deduction *}
wenzelm@44131
   687
wenzelm@44131
   688
text {* Isabelle can represent sequents directly, as in the
wenzelm@44131
   689
  object-logic LK.  But natural deduction is easier to work with, and
wenzelm@44131
   690
  most object-logics employ it.  Fortunately, we can simulate the
wenzelm@44131
   691
  sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} by the Isabelle formula
wenzelm@44131
   692
  @{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
wenzelm@44131
   693
  the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary.
wenzelm@44131
   694
  Elim-resolution plays a key role in simulating sequent proofs.
wenzelm@44131
   695
wenzelm@44131
   696
  We can easily handle reasoning on the left.  Elim-resolution with
wenzelm@44131
   697
  the rules @{text "(\<or>E)"}, @{text "(\<bottom>E)"} and @{text "(\<exists>E)"} achieves
wenzelm@44131
   698
  a similar effect as the corresponding sequent rules.  For the other
wenzelm@44131
   699
  connectives, we use sequent-style elimination rules instead of
wenzelm@44131
   700
  destruction rules such as @{text "(\<and>E1, 2)"} and @{text "(\<forall>E)"}.
wenzelm@44131
   701
  But note that the rule @{text "(\<not>L)"} has no effect under our
wenzelm@44131
   702
  representation of sequents!
wenzelm@44131
   703
  \[
wenzelm@44131
   704
  \infer[@{text "(\<not>L)"}]{@{text "\<not> P, \<Gamma> \<turnstile> \<Delta>"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P"}}
wenzelm@44131
   705
  \]
wenzelm@44131
   706
wenzelm@44131
   707
  What about reasoning on the right?  Introduction rules can only
wenzelm@44131
   708
  affect the formula in the conclusion, namely @{text "Q\<^sub>1"}.  The
wenzelm@44131
   709
  other right-side formulae are represented as negated assumptions,
wenzelm@44131
   710
  @{text "\<not> Q\<^sub>2, \<dots>, \<not> Q\<^sub>n"}.  In order to operate on one of these, it
wenzelm@44131
   711
  must first be exchanged with @{text "Q\<^sub>1"}.  Elim-resolution with the
wenzelm@44131
   712
  @{text swap} rule has this effect: @{text "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"}
wenzelm@44131
   713
wenzelm@44131
   714
  To ensure that swaps occur only when necessary, each introduction
wenzelm@44131
   715
  rule is converted into a swapped form: it is resolved with the
wenzelm@44131
   716
  second premise of @{text "(swap)"}.  The swapped form of @{text
wenzelm@44131
   717
  "(\<and>I)"}, which might be called @{text "(\<not>\<and>E)"}, is
wenzelm@44131
   718
  @{text [display] "\<not> (P \<and> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (\<not> R \<Longrightarrow> Q) \<Longrightarrow> R"}
wenzelm@44131
   719
wenzelm@44131
   720
  Similarly, the swapped form of @{text "(\<longrightarrow>I)"} is
wenzelm@44131
   721
  @{text [display] "\<not> (P \<longrightarrow> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P \<Longrightarrow> Q) \<Longrightarrow> R"}
wenzelm@44131
   722
wenzelm@44131
   723
  Swapped introduction rules are applied using elim-resolution, which
wenzelm@44131
   724
  deletes the negated formula.  Our representation of sequents also
wenzelm@44131
   725
  requires the use of ordinary introduction rules.  If we had no
wenzelm@44131
   726
  regard for readability of intermediate goal states, we could treat
wenzelm@44131
   727
  the right side more uniformly by representing sequents as @{text
wenzelm@44131
   728
  [display] "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> \<bottom>"}
wenzelm@44131
   729
*}
wenzelm@44131
   730
wenzelm@44131
   731
wenzelm@44131
   732
subsubsection {* Extra rules for the sequent calculus *}
wenzelm@44131
   733
wenzelm@44131
   734
text {* As mentioned, destruction rules such as @{text "(\<and>E1, 2)"} and
wenzelm@44131
   735
  @{text "(\<forall>E)"} must be replaced by sequent-style elimination rules.
wenzelm@44131
   736
  In addition, we need rules to embody the classical equivalence
wenzelm@44131
   737
  between @{text "P \<longrightarrow> Q"} and @{text "\<not> P \<or> Q"}.  The introduction
wenzelm@44131
   738
  rules @{text "(\<or>I1, 2)"} are replaced by a rule that simulates
wenzelm@44131
   739
  @{text "(\<or>R)"}: @{text [display] "(\<not> Q \<Longrightarrow> P) \<Longrightarrow> P \<or> Q"}
wenzelm@44131
   740
wenzelm@44131
   741
  The destruction rule @{text "(\<longrightarrow>E)"} is replaced by @{text [display]
wenzelm@44131
   742
  "(P \<longrightarrow> Q) \<Longrightarrow> (\<not> P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"}
wenzelm@44131
   743
wenzelm@44131
   744
  Quantifier replication also requires special rules.  In classical
wenzelm@44131
   745
  logic, @{text "\<exists>x. P x"} is equivalent to @{text "\<not> (\<forall>x. \<not> P x)"};
wenzelm@44131
   746
  the rules @{text "(\<exists>R)"} and @{text "(\<forall>L)"} are dual:
wenzelm@44131
   747
  \[
wenzelm@44131
   748
  \infer[@{text "(\<exists>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x"}}{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x, P t"}}
wenzelm@44131
   749
  \qquad
wenzelm@44131
   750
  \infer[@{text "(\<forall>L)"}]{@{text "\<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}{@{text "P t, \<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}
wenzelm@44131
   751
  \]
wenzelm@44131
   752
  Thus both kinds of quantifier may be replicated.  Theorems requiring
wenzelm@44131
   753
  multiple uses of a universal formula are easy to invent; consider
wenzelm@44131
   754
  @{text [display] "(\<forall>x. P x \<longrightarrow> P (f x)) \<and> P a \<longrightarrow> P (f\<^sup>n a)"} for any
wenzelm@44131
   755
  @{text "n > 1"}.  Natural examples of the multiple use of an
wenzelm@44131
   756
  existential formula are rare; a standard one is @{text "\<exists>x. \<forall>y. P x
wenzelm@44131
   757
  \<longrightarrow> P y"}.
wenzelm@44131
   758
wenzelm@44131
   759
  Forgoing quantifier replication loses completeness, but gains
wenzelm@44131
   760
  decidability, since the search space becomes finite.  Many useful
wenzelm@44131
   761
  theorems can be proved without replication, and the search generally
wenzelm@44131
   762
  delivers its verdict in a reasonable time.  To adopt this approach,
wenzelm@44131
   763
  represent the sequent rules @{text "(\<exists>R)"}, @{text "(\<exists>L)"} and
wenzelm@44131
   764
  @{text "(\<forall>R)"} by @{text "(\<exists>I)"}, @{text "(\<exists>E)"} and @{text "(\<forall>I)"},
wenzelm@44131
   765
  respectively, and put @{text "(\<forall>E)"} into elimination form: @{text
wenzelm@44131
   766
  [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"}
wenzelm@44131
   767
wenzelm@44131
   768
  Elim-resolution with this rule will delete the universal formula
wenzelm@44131
   769
  after a single use.  To replicate universal quantifiers, replace the
wenzelm@44131
   770
  rule by @{text [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"}
wenzelm@44131
   771
wenzelm@44131
   772
  To replicate existential quantifiers, replace @{text "(\<exists>I)"} by
wenzelm@44131
   773
  @{text [display] "(\<not> (\<exists>x. P x) \<Longrightarrow> P t) \<Longrightarrow> \<exists>x. P x"}
wenzelm@44131
   774
wenzelm@44131
   775
  All introduction rules mentioned above are also useful in swapped
wenzelm@44131
   776
  form.
wenzelm@44131
   777
wenzelm@44131
   778
  Replication makes the search space infinite; we must apply the rules
wenzelm@44131
   779
  with care.  The classical reasoner distinguishes between safe and
wenzelm@44131
   780
  unsafe rules, applying the latter only when there is no alternative.
wenzelm@44131
   781
  Depth-first search may well go down a blind alley; best-first search
wenzelm@44131
   782
  is better behaved in an infinite search space.  However, quantifier
wenzelm@44131
   783
  replication is too expensive to prove any but the simplest theorems.
wenzelm@44131
   784
*}
wenzelm@44131
   785
wenzelm@44131
   786
wenzelm@44132
   787
subsection {* Rule declarations *}
wenzelm@44132
   788
wenzelm@44132
   789
text {* The proof tools of the Classical Reasoner depend on
wenzelm@44132
   790
  collections of rules declared in the context, which are classified
wenzelm@44132
   791
  as introduction, elimination or destruction and as \emph{safe} or
wenzelm@44132
   792
  \emph{unsafe}.  In general, safe rules can be attempted blindly,
wenzelm@44132
   793
  while unsafe rules must be used with care.  A safe rule must never
wenzelm@44132
   794
  reduce a provable goal to an unprovable set of subgoals.
wenzelm@44132
   795
wenzelm@44132
   796
  The rule @{text "P \<Longrightarrow> P \<or> Q"} is unsafe because it reduces @{text "P
wenzelm@44132
   797
  \<or> Q"} to @{text "P"}, which might turn out as premature choice of an
wenzelm@44132
   798
  unprovable subgoal.  Any rule is unsafe whose premises contain new
wenzelm@44132
   799
  unknowns.  The elimination rule @{text "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"} is
wenzelm@44132
   800
  unsafe, since it is applied via elim-resolution, which discards the
wenzelm@44132
   801
  assumption @{text "\<forall>x. P x"} and replaces it by the weaker
wenzelm@44132
   802
  assumption @{text "P t"}.  The rule @{text "P t \<Longrightarrow> \<exists>x. P x"} is
wenzelm@44132
   803
  unsafe for similar reasons.  The quantifier duplication rule @{text
wenzelm@44132
   804
  "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"} is unsafe in a different sense:
wenzelm@44132
   805
  since it keeps the assumption @{text "\<forall>x. P x"}, it is prone to
wenzelm@44132
   806
  looping.  In classical first-order logic, all rules are safe except
wenzelm@44132
   807
  those mentioned above.
wenzelm@44132
   808
wenzelm@44132
   809
  The safe~/ unsafe distinction is vague, and may be regarded merely
wenzelm@44132
   810
  as a way of giving some rules priority over others.  One could argue
wenzelm@44132
   811
  that @{text "(\<or>E)"} is unsafe, because repeated application of it
wenzelm@44132
   812
  could generate exponentially many subgoals.  Induction rules are
wenzelm@44132
   813
  unsafe because inductive proofs are difficult to set up
wenzelm@44132
   814
  automatically.  Any inference is unsafe that instantiates an unknown
wenzelm@44132
   815
  in the proof state --- thus matching must be used, rather than
wenzelm@44132
   816
  unification.  Even proof by assumption is unsafe if it instantiates
wenzelm@44132
   817
  unknowns shared with other subgoals.
wenzelm@44132
   818
wenzelm@44132
   819
  \begin{matharray}{rcl}
wenzelm@44132
   820
    @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
wenzelm@44132
   821
    @{attribute_def intro} & : & @{text attribute} \\
wenzelm@44132
   822
    @{attribute_def elim} & : & @{text attribute} \\
wenzelm@44132
   823
    @{attribute_def dest} & : & @{text attribute} \\
wenzelm@44132
   824
    @{attribute_def rule} & : & @{text attribute} \\
wenzelm@44132
   825
    @{attribute_def iff} & : & @{text attribute} \\
wenzelm@44132
   826
    @{attribute_def swapped} & : & @{text attribute} \\
wenzelm@44132
   827
  \end{matharray}
wenzelm@44132
   828
wenzelm@44132
   829
  @{rail "
wenzelm@44132
   830
    (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}?
wenzelm@44132
   831
    ;
wenzelm@44132
   832
    @@{attribute rule} 'del'
wenzelm@44132
   833
    ;
wenzelm@44132
   834
    @@{attribute iff} (((() | 'add') '?'?) | 'del')
wenzelm@44132
   835
  "}
wenzelm@44132
   836
wenzelm@44132
   837
  \begin{description}
wenzelm@44132
   838
wenzelm@44132
   839
  \item @{command "print_claset"} prints the collection of rules
wenzelm@44132
   840
  declared to the Classical Reasoner, i.e.\ the @{ML_type claset}
wenzelm@44132
   841
  within the context.
wenzelm@44132
   842
wenzelm@44132
   843
  \item @{attribute intro}, @{attribute elim}, and @{attribute dest}
wenzelm@44132
   844
  declare introduction, elimination, and destruction rules,
wenzelm@44132
   845
  respectively.  By default, rules are considered as \emph{unsafe}
wenzelm@44132
   846
  (i.e.\ not applied blindly without backtracking), while ``@{text
wenzelm@44132
   847
  "!"}'' classifies as \emph{safe}.  Rule declarations marked by
wenzelm@44132
   848
  ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\
wenzelm@44132
   849
  \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps
wenzelm@44132
   850
  of the @{method rule} method).  The optional natural number
wenzelm@44132
   851
  specifies an explicit weight argument, which is ignored by the
wenzelm@44132
   852
  automated reasoning tools, but determines the search order of single
wenzelm@44132
   853
  rule steps.
wenzelm@44132
   854
wenzelm@44132
   855
  Introduction rules are those that can be applied using ordinary
wenzelm@44132
   856
  resolution.  Their swapped forms are generated internally, which
wenzelm@44132
   857
  will be applied using elim-resolution.  Elimination rules are
wenzelm@44132
   858
  applied using elim-resolution.  Rules are sorted by the number of
wenzelm@44132
   859
  new subgoals they will yield; rules that generate the fewest
wenzelm@44132
   860
  subgoals will be tried first.  Otherwise, later declarations take
wenzelm@44132
   861
  precedence over earlier ones.
wenzelm@44132
   862
wenzelm@44132
   863
  Rules already present in the context with the same classification
wenzelm@44132
   864
  are ignored.  A warning is printed if the rule has already been
wenzelm@44132
   865
  added with some other classification, but the rule is added anyway
wenzelm@44132
   866
  as requested.
wenzelm@44132
   867
wenzelm@44132
   868
  \item @{attribute rule}~@{text del} deletes all occurrences of a
wenzelm@44132
   869
  rule from the classical context, regardless of its classification as
wenzelm@44132
   870
  introduction~/ elimination~/ destruction and safe~/ unsafe.
wenzelm@44132
   871
wenzelm@44132
   872
  \item @{attribute iff} declares logical equivalences to the
wenzelm@44132
   873
  Simplifier and the Classical reasoner at the same time.
wenzelm@44132
   874
  Non-conditional rules result in a safe introduction and elimination
wenzelm@44132
   875
  pair; conditional ones are considered unsafe.  Rules with negative
wenzelm@44132
   876
  conclusion are automatically inverted (using @{text "\<not>"}-elimination
wenzelm@44132
   877
  internally).
wenzelm@44132
   878
wenzelm@44132
   879
  The ``@{text "?"}'' version of @{attribute iff} declares rules to
wenzelm@44132
   880
  the Isabelle/Pure context only, and omits the Simplifier
wenzelm@44132
   881
  declaration.
wenzelm@44132
   882
wenzelm@44132
   883
  \item @{attribute swapped} turns an introduction rule into an
wenzelm@44132
   884
  elimination, by resolving with the classical swap principle @{text
wenzelm@44132
   885
  "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"} in the second position.  This is mainly for
wenzelm@44132
   886
  illustrative purposes: the Classical Reasoner already swaps rules
wenzelm@44132
   887
  internally as explained above.
wenzelm@44132
   888
wenzelm@28760
   889
  \end{description}
wenzelm@26782
   890
*}
wenzelm@26782
   891
wenzelm@26782
   892
wenzelm@44236
   893
subsection {* Structured methods *}
wenzelm@44236
   894
wenzelm@44236
   895
text {*
wenzelm@44236
   896
  \begin{matharray}{rcl}
wenzelm@44236
   897
    @{method_def rule} & : & @{text method} \\
wenzelm@44236
   898
    @{method_def contradiction} & : & @{text method} \\
wenzelm@44236
   899
  \end{matharray}
wenzelm@44236
   900
wenzelm@44236
   901
  @{rail "
wenzelm@44236
   902
    @@{method rule} @{syntax thmrefs}?
wenzelm@44236
   903
  "}
wenzelm@44236
   904
wenzelm@44236
   905
  \begin{description}
wenzelm@44236
   906
wenzelm@44236
   907
  \item @{method rule} as offered by the Classical Reasoner is a
wenzelm@44236
   908
  refinement over the Pure one (see \secref{sec:pure-meth-att}).  Both
wenzelm@44236
   909
  versions work the same, but the classical version observes the
wenzelm@44236
   910
  classical rule context in addition to that of Isabelle/Pure.
wenzelm@44236
   911
wenzelm@44236
   912
  Common object logics (HOL, ZF, etc.) declare a rich collection of
wenzelm@44236
   913
  classical rules (even if these would qualify as intuitionistic
wenzelm@44236
   914
  ones), but only few declarations to the rule context of
wenzelm@44236
   915
  Isabelle/Pure (\secref{sec:pure-meth-att}).
wenzelm@44236
   916
wenzelm@44236
   917
  \item @{method contradiction} solves some goal by contradiction,
wenzelm@44236
   918
  deriving any result from both @{text "\<not> A"} and @{text A}.  Chained
wenzelm@44236
   919
  facts, which are guaranteed to participate, may appear in either
wenzelm@44236
   920
  order.
wenzelm@44236
   921
wenzelm@44236
   922
  \end{description}
wenzelm@44236
   923
*}
wenzelm@44236
   924
wenzelm@44236
   925
wenzelm@27040
   926
subsection {* Automated methods *}
wenzelm@26782
   927
wenzelm@26782
   928
text {*
wenzelm@26782
   929
  \begin{matharray}{rcl}
wenzelm@28761
   930
    @{method_def blast} & : & @{text method} \\
wenzelm@44134
   931
    @{method_def auto} & : & @{text method} \\
wenzelm@44134
   932
    @{method_def force} & : & @{text method} \\
wenzelm@28761
   933
    @{method_def fast} & : & @{text method} \\
wenzelm@28761
   934
    @{method_def slow} & : & @{text method} \\
wenzelm@28761
   935
    @{method_def best} & : & @{text method} \\
wenzelm@28761
   936
    @{method_def fastsimp} & : & @{text method} \\
wenzelm@28761
   937
    @{method_def slowsimp} & : & @{text method} \\
wenzelm@28761
   938
    @{method_def bestsimp} & : & @{text method} \\
wenzelm@44238
   939
    @{method_def deepen} & : & @{text method} \\
wenzelm@26782
   940
  \end{matharray}
wenzelm@26782
   941
wenzelm@43467
   942
  @{rail "
wenzelm@44134
   943
    @@{method blast} @{syntax nat}? (@{syntax clamod} * )
wenzelm@44134
   944
    ;
wenzelm@43467
   945
    @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * )
wenzelm@26782
   946
    ;
wenzelm@44134
   947
    @@{method force} (@{syntax clasimpmod} * )
wenzelm@26782
   948
    ;
wenzelm@44134
   949
    (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * )
wenzelm@44134
   950
    ;
wenzelm@44134
   951
    (@@{method fastsimp} | @@{method slowsimp} | @@{method bestsimp})
wenzelm@44134
   952
      (@{syntax clasimpmod} * )
wenzelm@44134
   953
    ;
wenzelm@44238
   954
    @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * )
wenzelm@44238
   955
    ;
wenzelm@44134
   956
    @{syntax_def clamod}:
wenzelm@44134
   957
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs}
wenzelm@44134
   958
    ;
wenzelm@43467
   959
    @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') |
wenzelm@26782
   960
      ('cong' | 'split') (() | 'add' | 'del') |
wenzelm@26782
   961
      'iff' (((() | 'add') '?'?) | 'del') |
wenzelm@43467
   962
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs}
wenzelm@43467
   963
  "}
wenzelm@26782
   964
wenzelm@28760
   965
  \begin{description}
wenzelm@26782
   966
wenzelm@44134
   967
  \item @{method blast} is a separate classical tableau prover that
wenzelm@44134
   968
  uses the same classical rule declarations as explained before.
wenzelm@26782
   969
wenzelm@44134
   970
  Proof search is coded directly in ML using special data structures.
wenzelm@44134
   971
  A successful proof is then reconstructed using regular Isabelle
wenzelm@44134
   972
  inferences.  It is faster and more powerful than the other classical
wenzelm@44134
   973
  reasoning tools, but has major limitations too.
wenzelm@44134
   974
wenzelm@44134
   975
  \begin{itemize}
wenzelm@44134
   976
wenzelm@44134
   977
  \item It does not use the classical wrapper tacticals, such as the
wenzelm@44134
   978
  integration with the Simplifier of @{method fastsimp}.
wenzelm@44134
   979
wenzelm@44134
   980
  \item It does not perform higher-order unification, as needed by the
wenzelm@44134
   981
  rule @{thm [source=false] rangeI} in HOL.  There are often
wenzelm@44134
   982
  alternatives to such rules, for example @{thm [source=false]
wenzelm@44134
   983
  range_eqI}.
wenzelm@44134
   984
wenzelm@44134
   985
  \item Function variables may only be applied to parameters of the
wenzelm@44134
   986
  subgoal.  (This restriction arises because the prover does not use
wenzelm@44134
   987
  higher-order unification.)  If other function variables are present
wenzelm@44134
   988
  then the prover will fail with the message \texttt{Function Var's
wenzelm@44134
   989
  argument not a bound variable}.
wenzelm@44134
   990
wenzelm@44134
   991
  \item Its proof strategy is more general than @{method fast} but can
wenzelm@44134
   992
  be slower.  If @{method blast} fails or seems to be running forever,
wenzelm@44134
   993
  try @{method fast} and the other proof tools described below.
wenzelm@44134
   994
wenzelm@44134
   995
  \end{itemize}
wenzelm@44134
   996
wenzelm@44134
   997
  The optional integer argument specifies a bound for the number of
wenzelm@44134
   998
  unsafe steps used in a proof.  By default, @{method blast} starts
wenzelm@44134
   999
  with a bound of 0 and increases it successively to 20.  In contrast,
wenzelm@44134
  1000
  @{text "(blast lim)"} tries to prove the goal using a search bound
wenzelm@44134
  1001
  of @{text "lim"}.  Sometimes a slow proof using @{method blast} can
wenzelm@44134
  1002
  be made much faster by supplying the successful search bound to this
wenzelm@44134
  1003
  proof method instead.
wenzelm@44134
  1004
wenzelm@44134
  1005
  \item @{method auto} combines classical reasoning with
wenzelm@44134
  1006
  simplification.  It is intended for situations where there are a lot
wenzelm@44134
  1007
  of mostly trivial subgoals; it proves all the easy ones, leaving the
wenzelm@44134
  1008
  ones it cannot prove.  Occasionally, attempting to prove the hard
wenzelm@44134
  1009
  ones may take a long time.
wenzelm@44134
  1010
wenzelm@44214
  1011
  The optional depth arguments in @{text "(auto m n)"} refer to its
wenzelm@44214
  1012
  builtin classical reasoning procedures: @{text m} (default 4) is for
wenzelm@44214
  1013
  @{method blast}, which is tried first, and @{text n} (default 2) is
wenzelm@44214
  1014
  for a slower but more general alternative that also takes wrappers
wenzelm@44214
  1015
  into account.
wenzelm@44134
  1016
wenzelm@44134
  1017
  \item @{method force} is intended to prove the first subgoal
wenzelm@44134
  1018
  completely, using many fancy proof tools and performing a rather
wenzelm@44134
  1019
  exhaustive search.  As a result, proof attempts may take rather long
wenzelm@44134
  1020
  or diverge easily.
wenzelm@44134
  1021
wenzelm@44134
  1022
  \item @{method fast}, @{method best}, @{method slow} attempt to
wenzelm@44134
  1023
  prove the first subgoal using sequent-style reasoning as explained
wenzelm@44134
  1024
  before.  Unlike @{method blast}, they construct proofs directly in
wenzelm@44134
  1025
  Isabelle.
wenzelm@44134
  1026
wenzelm@44134
  1027
  There is a difference in search strategy and back-tracking: @{method
wenzelm@44134
  1028
  fast} uses depth-first search and @{method best} uses best-first
wenzelm@44134
  1029
  search (guided by a heuristic function: normally the total size of
wenzelm@44134
  1030
  the proof state).
wenzelm@44134
  1031
wenzelm@44134
  1032
  Method @{method slow} is like @{method fast}, but conducts a broader
wenzelm@44134
  1033
  search: it may, when backtracking from a failed proof attempt, undo
wenzelm@44134
  1034
  even the step of proving a subgoal by assumption.
wenzelm@44134
  1035
wenzelm@44134
  1036
  \item @{method fastsimp}, @{method slowsimp}, @{method bestsimp} are
wenzelm@44134
  1037
  like @{method fast}, @{method slow}, @{method best}, respectively,
wenzelm@44134
  1038
  but use the Simplifier as additional wrapper.
wenzelm@44134
  1039
wenzelm@44238
  1040
  \item @{method deepen} works by exhaustive search up to a certain
wenzelm@44238
  1041
  depth.  The start depth is 4 (unless specified explicitly), and the
wenzelm@44238
  1042
  depth is increased iteratively up to 10.  Unsafe rules are modified
wenzelm@44238
  1043
  to preserve the formula they act on, so that it be used repeatedly.
wenzelm@44238
  1044
  This method can prove more goals than @{method fast}, but is much
wenzelm@44238
  1045
  slower, for example if the assumptions have many universal
wenzelm@44238
  1046
  quantifiers.
wenzelm@44238
  1047
wenzelm@44134
  1048
  \end{description}
wenzelm@44134
  1049
wenzelm@44134
  1050
  Any of the above methods support additional modifiers of the context
wenzelm@44134
  1051
  of classical (and simplifier) rules, but the ones related to the
wenzelm@44134
  1052
  Simplifier are explicitly prefixed by @{text simp} here.  The
wenzelm@44134
  1053
  semantics of these ad-hoc rule declarations is analogous to the
wenzelm@44134
  1054
  attributes given before.  Facts provided by forward chaining are
wenzelm@44134
  1055
  inserted into the goal before commencing proof search.
wenzelm@44134
  1056
*}
wenzelm@44134
  1057
wenzelm@44134
  1058
wenzelm@44134
  1059
subsection {* Semi-automated methods *}
wenzelm@44134
  1060
wenzelm@44134
  1061
text {* These proof methods may help in situations when the
wenzelm@44134
  1062
  fully-automated tools fail.  The result is a simpler subgoal that
wenzelm@44134
  1063
  can be tackled by other means, such as by manual instantiation of
wenzelm@44134
  1064
  quantifiers.
wenzelm@44134
  1065
wenzelm@44134
  1066
  \begin{matharray}{rcl}
wenzelm@44134
  1067
    @{method_def safe} & : & @{text method} \\
wenzelm@44134
  1068
    @{method_def clarify} & : & @{text method} \\
wenzelm@44134
  1069
    @{method_def clarsimp} & : & @{text method} \\
wenzelm@44134
  1070
  \end{matharray}
wenzelm@44134
  1071
wenzelm@44134
  1072
  @{rail "
wenzelm@44134
  1073
    (@@{method safe} | @@{method clarify}) (@{syntax clamod} * )
wenzelm@44134
  1074
    ;
wenzelm@44134
  1075
    @@{method clarsimp} (@{syntax clasimpmod} * )
wenzelm@44134
  1076
  "}
wenzelm@44134
  1077
wenzelm@44134
  1078
  \begin{description}
wenzelm@44134
  1079
wenzelm@44134
  1080
  \item @{method safe} repeatedly performs safe steps on all subgoals.
wenzelm@44134
  1081
  It is deterministic, with at most one outcome.
wenzelm@44134
  1082
wenzelm@44237
  1083
  \item @{method clarify} performs a series of safe steps without
wenzelm@44237
  1084
  splitting subgoals; see also @{ML clarify_step_tac}.
wenzelm@44134
  1085
wenzelm@44134
  1086
  \item @{method clarsimp} acts like @{method clarify}, but also does
wenzelm@44134
  1087
  simplification.  Note that if the Simplifier context includes a
wenzelm@44134
  1088
  splitter for the premises, the subgoal may still be split.
wenzelm@26782
  1089
wenzelm@28760
  1090
  \end{description}
wenzelm@26782
  1091
*}
wenzelm@26782
  1092
wenzelm@26782
  1093
wenzelm@44237
  1094
subsection {* Single-step tactics *}
wenzelm@44237
  1095
wenzelm@44237
  1096
text {*
wenzelm@44237
  1097
  \begin{matharray}{rcl}
wenzelm@44237
  1098
    @{index_ML safe_step_tac: "Proof.context -> int -> tactic"} \\
wenzelm@44237
  1099
    @{index_ML inst_step_tac: "Proof.context -> int -> tactic"} \\
wenzelm@44237
  1100
    @{index_ML step_tac: "Proof.context -> int -> tactic"} \\
wenzelm@44237
  1101
    @{index_ML slow_step_tac: "Proof.context -> int -> tactic"} \\
wenzelm@44237
  1102
    @{index_ML clarify_step_tac: "Proof.context -> int -> tactic"} \\
wenzelm@44237
  1103
  \end{matharray}
wenzelm@44237
  1104
wenzelm@44237
  1105
  These are the primitive tactics behind the (semi)automated proof
wenzelm@44237
  1106
  methods of the Classical Reasoner.  By calling them yourself, you
wenzelm@44237
  1107
  can execute these procedures one step at a time.
wenzelm@44237
  1108
wenzelm@44237
  1109
  \begin{description}
wenzelm@44237
  1110
wenzelm@44237
  1111
  \item @{ML safe_step_tac}~@{text "ctxt i"} performs a safe step on
wenzelm@44237
  1112
  subgoal @{text i}.  The safe wrapper tacticals are applied to a
wenzelm@44237
  1113
  tactic that may include proof by assumption or Modus Ponens (taking
wenzelm@44237
  1114
  care not to instantiate unknowns), or substitution.
wenzelm@44237
  1115
wenzelm@44237
  1116
  \item @{ML inst_step_tac} is like @{ML safe_step_tac}, but allows
wenzelm@44237
  1117
  unknowns to be instantiated.
wenzelm@44237
  1118
wenzelm@44237
  1119
  \item @{ML step_tac}~@{text "ctxt i"} is the basic step of the proof
wenzelm@44237
  1120
  procedure.  The unsafe wrapper tacticals are applied to a tactic
wenzelm@44237
  1121
  that tries @{ML safe_tac}, @{ML inst_step_tac}, or applies an unsafe
wenzelm@44237
  1122
  rule from the context.
wenzelm@44237
  1123
wenzelm@44237
  1124
  \item @{ML slow_step_tac} resembles @{ML step_tac}, but allows
wenzelm@44237
  1125
  backtracking between using safe rules with instantiation (@{ML
wenzelm@44237
  1126
  inst_step_tac}) and using unsafe rules.  The resulting search space
wenzelm@44237
  1127
  is larger.
wenzelm@44237
  1128
wenzelm@44237
  1129
  \item @{ML clarify_step_tac}~@{text "ctxt i"} performs a safe step
wenzelm@44237
  1130
  on subgoal @{text i}.  No splitting step is applied; for example,
wenzelm@44237
  1131
  the subgoal @{text "A \<and> B"} is left as a conjunction.  Proof by
wenzelm@44237
  1132
  assumption, Modus Ponens, etc., may be performed provided they do
wenzelm@44237
  1133
  not instantiate unknowns.  Assumptions of the form @{text "x = t"}
wenzelm@44237
  1134
  may be eliminated.  The safe wrapper tactical is applied.
wenzelm@44237
  1135
wenzelm@44237
  1136
  \end{description}
wenzelm@44237
  1137
*}
wenzelm@44237
  1138
wenzelm@44237
  1139
wenzelm@27044
  1140
section {* Object-logic setup \label{sec:object-logic} *}
wenzelm@26790
  1141
wenzelm@26790
  1142
text {*
wenzelm@26790
  1143
  \begin{matharray}{rcl}
wenzelm@28761
  1144
    @{command_def "judgment"} & : & @{text "theory \<rightarrow> theory"} \\
wenzelm@28761
  1145
    @{method_def atomize} & : & @{text method} \\
wenzelm@28761
  1146
    @{attribute_def atomize} & : & @{text attribute} \\
wenzelm@28761
  1147
    @{attribute_def rule_format} & : & @{text attribute} \\
wenzelm@28761
  1148
    @{attribute_def rulify} & : & @{text attribute} \\
wenzelm@26790
  1149
  \end{matharray}
wenzelm@26790
  1150
wenzelm@26790
  1151
  The very starting point for any Isabelle object-logic is a ``truth
wenzelm@26790
  1152
  judgment'' that links object-level statements to the meta-logic
wenzelm@26790
  1153
  (with its minimal language of @{text prop} that covers universal
wenzelm@26790
  1154
  quantification @{text "\<And>"} and implication @{text "\<Longrightarrow>"}).
wenzelm@26790
  1155
wenzelm@26790
  1156
  Common object-logics are sufficiently expressive to internalize rule
wenzelm@26790
  1157
  statements over @{text "\<And>"} and @{text "\<Longrightarrow>"} within their own
wenzelm@26790
  1158
  language.  This is useful in certain situations where a rule needs
wenzelm@26790
  1159
  to be viewed as an atomic statement from the meta-level perspective,
wenzelm@26790
  1160
  e.g.\ @{text "\<And>x. x \<in> A \<Longrightarrow> P x"} versus @{text "\<forall>x \<in> A. P x"}.
wenzelm@26790
  1161
wenzelm@26790
  1162
  From the following language elements, only the @{method atomize}
wenzelm@26790
  1163
  method and @{attribute rule_format} attribute are occasionally
wenzelm@26790
  1164
  required by end-users, the rest is for those who need to setup their
wenzelm@26790
  1165
  own object-logic.  In the latter case existing formulations of
wenzelm@26790
  1166
  Isabelle/FOL or Isabelle/HOL may be taken as realistic examples.
wenzelm@26790
  1167
wenzelm@26790
  1168
  Generic tools may refer to the information provided by object-logic
wenzelm@26790
  1169
  declarations internally.
wenzelm@26790
  1170
wenzelm@43467
  1171
  @{rail "
wenzelm@43467
  1172
    @@{command judgment} @{syntax constdecl}
wenzelm@26790
  1173
    ;
wenzelm@43467
  1174
    @@{attribute atomize} ('(' 'full' ')')?
wenzelm@26790
  1175
    ;
wenzelm@43467
  1176
    @@{attribute rule_format} ('(' 'noasm' ')')?
wenzelm@43467
  1177
  "}
wenzelm@26790
  1178
wenzelm@28760
  1179
  \begin{description}
wenzelm@26790
  1180
  
wenzelm@28760
  1181
  \item @{command "judgment"}~@{text "c :: \<sigma> (mx)"} declares constant
wenzelm@28760
  1182
  @{text c} as the truth judgment of the current object-logic.  Its
wenzelm@28760
  1183
  type @{text \<sigma>} should specify a coercion of the category of
wenzelm@28760
  1184
  object-level propositions to @{text prop} of the Pure meta-logic;
wenzelm@28760
  1185
  the mixfix annotation @{text "(mx)"} would typically just link the
wenzelm@28760
  1186
  object language (internally of syntactic category @{text logic})
wenzelm@28760
  1187
  with that of @{text prop}.  Only one @{command "judgment"}
wenzelm@28760
  1188
  declaration may be given in any theory development.
wenzelm@26790
  1189
  
wenzelm@28760
  1190
  \item @{method atomize} (as a method) rewrites any non-atomic
wenzelm@26790
  1191
  premises of a sub-goal, using the meta-level equations declared via
wenzelm@26790
  1192
  @{attribute atomize} (as an attribute) beforehand.  As a result,
wenzelm@26790
  1193
  heavily nested goals become amenable to fundamental operations such
wenzelm@43497
  1194
  as resolution (cf.\ the @{method (Pure) rule} method).  Giving the ``@{text
wenzelm@26790
  1195
  "(full)"}'' option here means to turn the whole subgoal into an
wenzelm@26790
  1196
  object-statement (if possible), including the outermost parameters
wenzelm@26790
  1197
  and assumptions as well.
wenzelm@26790
  1198
wenzelm@26790
  1199
  A typical collection of @{attribute atomize} rules for a particular
wenzelm@26790
  1200
  object-logic would provide an internalization for each of the
wenzelm@26790
  1201
  connectives of @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"}.
wenzelm@26790
  1202
  Meta-level conjunction should be covered as well (this is
wenzelm@26790
  1203
  particularly important for locales, see \secref{sec:locale}).
wenzelm@26790
  1204
wenzelm@28760
  1205
  \item @{attribute rule_format} rewrites a theorem by the equalities
wenzelm@28760
  1206
  declared as @{attribute rulify} rules in the current object-logic.
wenzelm@28760
  1207
  By default, the result is fully normalized, including assumptions
wenzelm@28760
  1208
  and conclusions at any depth.  The @{text "(no_asm)"} option
wenzelm@28760
  1209
  restricts the transformation to the conclusion of a rule.
wenzelm@26790
  1210
wenzelm@26790
  1211
  In common object-logics (HOL, FOL, ZF), the effect of @{attribute
wenzelm@26790
  1212
  rule_format} is to replace (bounded) universal quantification
wenzelm@26790
  1213
  (@{text "\<forall>"}) and implication (@{text "\<longrightarrow>"}) by the corresponding
wenzelm@26790
  1214
  rule statements over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
wenzelm@26790
  1215
wenzelm@28760
  1216
  \end{description}
wenzelm@26790
  1217
*}
wenzelm@26790
  1218
wenzelm@26782
  1219
end