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