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