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