src/Doc/Isar_Ref/Generic.thy
author wenzelm
Mon, 15 Jun 2015 14:10:41 +0200
changeset 60484 98ee86354354
parent 59921 5b919b13feee
child 60554 c0e1c121c7c0
permissions -rw-r--r--
moved sections;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     1
theory Generic
42651
e3fdb7c96be5 formal Base theory;
wenzelm
parents: 42626
diff changeset
     2
imports Base Main
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     3
begin
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     4
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     5
chapter \<open>Generic tools and packages \label{ch:gen-tools}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     7
section \<open>Configuration options \label{sec:config}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
     8
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     9
text \<open>Isabelle/Pure maintains a record of named configuration
40291
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    10
  options within the theory or proof context, with values of type
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    11
  @{ML_type bool}, @{ML_type int}, @{ML_type real}, or @{ML_type
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    12
  string}.  Tools may declare options in ML, and then refer to these
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    13
  values (relative to the context).  Thus global reference variables
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    14
  are easily avoided.  The user may change the value of a
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    15
  configuration option by means of an associated attribute of the same
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40255
diff changeset
    16
  name.  This form of context declaration works particularly well with
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    17
  commands such as @{command "declare"} or @{command "using"} like
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    18
  this:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    19
\<close>
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    20
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
    21
(*<*)experiment begin(*>*)
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    22
declare [[show_main_goal = false]]
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    23
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    24
notepad
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    25
begin
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    26
  note [[show_main_goal = true]]
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    27
end
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
    28
(*<*)end(*>*)
42655
eb95e2f3b218 updated configuration options -- no ML here;
wenzelm
parents: 42651
diff changeset
    29
59921
5b919b13feee obsolete (see 8b7caf447357);
wenzelm
parents: 59917
diff changeset
    30
text \<open>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    31
  \begin{matharray}{rcll}
52060
179236c82c2a renamed 'print_configs' to 'print_options';
wenzelm
parents: 52037
diff changeset
    32
    @{command_def "print_options"} & : & @{text "context \<rightarrow>"} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    33
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    34
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
    35
  @{rail \<open>
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
    36
    @@{command print_options} ('!'?)
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
    37
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    38
    @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
    39
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    40
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    41
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    42
  
52060
179236c82c2a renamed 'print_configs' to 'print_options';
wenzelm
parents: 52037
diff changeset
    43
  \item @{command "print_options"} prints the available configuration
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
    44
  options, with names, types, and current values; the ``@{text "!"}'' option
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
    45
  indicates extra verbosity.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    46
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    47
  \item @{text "name = value"} as an attribute expression modifies the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    48
  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
    49
  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
    50
  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
    51
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    52
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    53
\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    54
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    55
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    56
section \<open>Basic proof tools\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    57
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    58
subsection \<open>Miscellaneous methods and attributes \label{sec:misc-meth-att}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    59
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    60
text \<open>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    61
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    62
    @{method_def unfold} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    63
    @{method_def fold} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    64
    @{method_def insert} & : & @{text method} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    65
    @{method_def erule}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    66
    @{method_def drule}@{text "\<^sup>*"} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    67
    @{method_def frule}@{text "\<^sup>*"} & : & @{text method} \\
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
    68
    @{method_def intro} & : & @{text method} \\
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
    69
    @{method_def elim} & : & @{text method} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    70
    @{method_def succeed} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
    71
    @{method_def fail} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    72
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    73
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
    74
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    75
    (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thmrefs}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    76
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    77
    (@@{method erule} | @@{method drule} | @@{method frule})
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
    78
      ('(' @{syntax nat} ')')? @{syntax thmrefs}
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
    79
    ;
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
    80
    (@@{method intro} | @@{method elim}) @{syntax thmrefs}?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
    81
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    82
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    83
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    84
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    85
  \item @{method 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
    86
  "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
    87
  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
    88
  subject to rewriting as well.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    89
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    90
  \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
    91
  into all goals of the proof state.  Note that current facts
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
    92
  indicated for forward chaining are ignored.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
    93
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    94
  \item @{method erule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, @{method
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    95
  drule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, and @{method frule}~@{text
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    96
  "a\<^sub>1 \<dots> a\<^sub>n"} are similar to the basic @{method rule}
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    97
  method (see \secref{sec:pure-meth-att}), but apply rules by
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
    98
  elim-resolution, destruct-resolution, and forward-resolution,
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58310
diff changeset
    99
  respectively @{cite "isabelle-implementation"}.  The optional natural
30397
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   100
  number argument (default 0) specifies additional assumption steps to
b6212ae21656 markup antiquotation options;
wenzelm
parents: 30168
diff changeset
   101
  be performed here.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   102
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   103
  Note that these methods are improper ones, mainly serving for
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   104
  experimentation and tactic script emulation.  Different modes of
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   105
  basic rule application are usually expressed in Isar at the proof
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   106
  language level, rather than via implicit proof state manipulations.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   107
  For example, a proper single-step elimination would be done using
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   108
  the plain @{method rule} method, with forward chaining of current
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   109
  facts.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   110
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
   111
  \item @{method intro} and @{method elim} repeatedly refine some goal
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
   112
  by intro- or elim-resolution, after having inserted any chained
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
   113
  facts.  Exactly the rules given as arguments are taken into account;
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
   114
  this allows fine-tuned decomposition of a proof problem, in contrast
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
   115
  to common automated tools.
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
   116
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   117
  \item @{method succeed} yields a single (unchanged) result; it is
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   118
  the identity of the ``@{text ","}'' method combinator (cf.\
28754
6f2e67a3dfaa moved section "Proof method expressions" to proof chapter;
wenzelm
parents: 27248
diff changeset
   119
  \secref{sec:proof-meth}).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   120
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   121
  \item @{method fail} yields an empty result sequence; it is the
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   122
  identity of the ``@{text "|"}'' method combinator (cf.\
28754
6f2e67a3dfaa moved section "Proof method expressions" to proof chapter;
wenzelm
parents: 27248
diff changeset
   123
  \secref{sec:proof-meth}).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   124
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   125
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   126
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   127
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   128
    @{attribute_def tagged} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   129
    @{attribute_def untagged} & : & @{text attribute} \\[0.5ex]
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   130
    @{attribute_def THEN} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   131
    @{attribute_def unfolded} & : & @{text attribute} \\
47497
c78c6e1ec75d document attribute "abs_def";
wenzelm
parents: 46706
diff changeset
   132
    @{attribute_def folded} & : & @{text attribute} \\
c78c6e1ec75d document attribute "abs_def";
wenzelm
parents: 46706
diff changeset
   133
    @{attribute_def abs_def} & : & @{text attribute} \\[0.5ex]
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   134
    @{attribute_def rotated} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   135
    @{attribute_def (Pure) elim_format} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   136
    @{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   137
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   138
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   139
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   140
    @@{attribute tagged} @{syntax name} @{syntax name}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   141
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   142
    @@{attribute untagged} @{syntax name}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   143
    ;
48205
09c2a3d9aa22 discontinued obsolete attribute "COMP";
wenzelm
parents: 47967
diff changeset
   144
    @@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thmref}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   145
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   146
    (@@{attribute unfolded} | @@{attribute folded}) @{syntax thmrefs}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   147
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   148
    @@{attribute rotated} @{syntax int}?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   149
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   150
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   151
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   152
28764
b65194fe4434 fixed/tuned syntax for attribute "tagged";
wenzelm
parents: 28761
diff changeset
   153
  \item @{attribute tagged}~@{text "name value"} and @{attribute
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   154
  untagged}~@{text name} add and remove \emph{tags} of some theorem.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   155
  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
   156
  The first string is considered the tag name, the second its value.
b65194fe4434 fixed/tuned syntax for attribute "tagged";
wenzelm
parents: 28761
diff changeset
   157
  Note that @{attribute untagged} removes any tags of the same name.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   158
48205
09c2a3d9aa22 discontinued obsolete attribute "COMP";
wenzelm
parents: 47967
diff changeset
   159
  \item @{attribute THEN}~@{text a} composes rules by resolution; it
09c2a3d9aa22 discontinued obsolete attribute "COMP";
wenzelm
parents: 47967
diff changeset
   160
  resolves with the first premise of @{text a} (an alternative
09c2a3d9aa22 discontinued obsolete attribute "COMP";
wenzelm
parents: 47967
diff changeset
   161
  position may be also specified).  See also @{ML_op "RS"} in
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58310
diff changeset
   162
  @{cite "isabelle-implementation"}.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   163
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   164
  \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
   165
  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
   166
  definitions throughout a rule.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   167
47497
c78c6e1ec75d document attribute "abs_def";
wenzelm
parents: 46706
diff changeset
   168
  \item @{attribute abs_def} turns an equation of the form @{prop "f x
c78c6e1ec75d document attribute "abs_def";
wenzelm
parents: 46706
diff changeset
   169
  y \<equiv> t"} into @{prop "f \<equiv> \<lambda>x y. t"}, which ensures that @{method
c78c6e1ec75d document attribute "abs_def";
wenzelm
parents: 46706
diff changeset
   170
  simp} or @{method unfold} steps always expand it.  This also works
c78c6e1ec75d document attribute "abs_def";
wenzelm
parents: 46706
diff changeset
   171
  for object-logic equality.
c78c6e1ec75d document attribute "abs_def";
wenzelm
parents: 46706
diff changeset
   172
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   173
  \item @{attribute rotated}~@{text n} rotate the premises of a
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   174
  theorem by @{text n} (default 1).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   175
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   176
  \item @{attribute (Pure) elim_format} turns a destruction rule into
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   177
  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
   178
  (PROP A \<Longrightarrow> PROP B) \<Longrightarrow> PROP B"}.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   179
  
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   180
  Note that the Classical Reasoner (\secref{sec:classical}) provides
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   181
  its own version of this operation.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   182
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   183
  \item @{attribute no_vars} replaces schematic variables by free
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   184
  ones; this is mainly for tuning output of pretty printed theorems.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   185
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   186
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   187
\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   188
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   189
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   190
subsection \<open>Low-level equational reasoning\<close>
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   191
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   192
text \<open>
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   193
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   194
    @{method_def subst} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   195
    @{method_def hypsubst} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   196
    @{method_def split} & : & @{text method} \\
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   197
  \end{matharray}
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   198
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   199
  @{rail \<open>
55029
61a6bf7d4b02 clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents: 52060
diff changeset
   200
    @@{method subst} ('(' 'asm' ')')? \<newline> ('(' (@{syntax nat}+) ')')? @{syntax thmref}
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   201
    ;
44094
f7bbfdf4b4a7 updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents: 43367
diff changeset
   202
    @@{method split} @{syntax thmrefs}
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   203
  \<close>}
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   204
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   205
  These methods provide low-level facilities for equational reasoning
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   206
  that are intended for specialized applications only.  Normally,
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   207
  single step calculations would be performed in a structured text
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   208
  (see also \secref{sec:calculation}), while the Simplifier methods
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   209
  provide the canonical way for automated normalization (see
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   210
  \secref{sec:simplifier}).
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   211
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   212
  \begin{description}
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   213
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   214
  \item @{method subst}~@{text eq} performs a single substitution step
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   215
  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
   216
  equality.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   217
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   218
  \item @{method subst}~@{text "(asm) eq"} substitutes in an
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   219
  assumption.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   220
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   221
  \item @{method subst}~@{text "(i \<dots> j) eq"} performs several
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   222
  substitutions in the conclusion. The numbers @{text i} to @{text j}
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   223
  indicate the positions to substitute at.  Positions are ordered from
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   224
  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
   225
  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
   226
  where commutativity of @{text "+"} is applicable: 1 refers to @{text
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   227
  "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
   228
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   229
  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
   230
  (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
   231
  assume all substitutions are performed simultaneously.  Otherwise
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   232
  the behaviour of @{text subst} is not specified.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   233
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   234
  \item @{method subst}~@{text "(asm) (i \<dots> j) eq"} performs the
27071
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   235
  substitutions in the assumptions. The positions refer to the
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   236
  assumptions in order from left to right.  For example, given in a
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   237
  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
   238
  commutativity of @{text "+"} is the subterm @{text "a + b"} and
614c045c5fd4 clarification of "subst" by Lucas Dixon;
wenzelm
parents: 27044
diff changeset
   239
  position 2 is the subterm @{text "c + d"}.
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   240
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   241
  \item @{method hypsubst} performs substitution using some
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   242
  assumption; this only works for equations of the form @{text "x =
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   243
  t"} where @{text x} is a free or bound variable.
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   244
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   245
  \item @{method split}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} performs single-step case
44094
f7bbfdf4b4a7 updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents: 43367
diff changeset
   246
  splitting using the given rules.  Splitting is performed in the
f7bbfdf4b4a7 updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents: 43367
diff changeset
   247
  conclusion or some assumption of the subgoal, depending of the
f7bbfdf4b4a7 updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents: 43367
diff changeset
   248
  structure of the rule.
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   249
  
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   250
  Note that the @{method simp} method already involves repeated
44094
f7bbfdf4b4a7 updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents: 43367
diff changeset
   251
  application of split rules as declared in the current context, using
f7bbfdf4b4a7 updated documentation of method "split" according to e6a4bb832b46;
wenzelm
parents: 43367
diff changeset
   252
  @{attribute split}, for example.
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   253
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   254
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   255
\<close>
27044
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   256
c4eaa7140532 moved subst/hypsubst to "Basic proof tools";
wenzelm
parents: 27040
diff changeset
   257
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   258
section \<open>The Simplifier \label{sec:simplifier}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   259
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   260
text \<open>The Simplifier performs conditional and unconditional
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   261
  rewriting and uses contextual information: rule declarations in the
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   262
  background theory or local proof context are taken into account, as
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   263
  well as chained facts and subgoal premises (``local assumptions'').
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   264
  There are several general hooks that allow to modify the
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   265
  simplification strategy, or incorporate other proof tools that solve
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   266
  sub-problems, produce rewrite rules on demand etc.
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   267
50075
ceb877c315a1 (re)moved old material about Simplifier;
wenzelm
parents: 50071
diff changeset
   268
  The rewriting strategy is always strictly bottom up, except for
ceb877c315a1 (re)moved old material about Simplifier;
wenzelm
parents: 50071
diff changeset
   269
  congruence rules, which are applied while descending into a term.
ceb877c315a1 (re)moved old material about Simplifier;
wenzelm
parents: 50071
diff changeset
   270
  Conditions in conditional rewrite rules are solved recursively
ceb877c315a1 (re)moved old material about Simplifier;
wenzelm
parents: 50071
diff changeset
   271
  before the rewrite rule is applied.
ceb877c315a1 (re)moved old material about Simplifier;
wenzelm
parents: 50071
diff changeset
   272
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   273
  The default Simplifier setup of major object logics (HOL, HOLCF,
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   274
  FOL, ZF) makes the Simplifier ready for immediate use, without
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   275
  engaging into the internal structures.  Thus it serves as
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   276
  general-purpose proof tool with the main focus on equational
50075
ceb877c315a1 (re)moved old material about Simplifier;
wenzelm
parents: 50071
diff changeset
   277
  reasoning, and a bit more than that.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   278
\<close>
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   279
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   280
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   281
subsection \<open>Simplification methods \label{sec:simp-meth}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   282
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   283
text \<open>
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   284
  \begin{tabular}{rcll}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   285
    @{method_def simp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   286
    @{method_def simp_all} & : & @{text method} \\
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   287
    @{attribute_def simp_depth_limit} & : & @{text attribute} & default @{text 100} \\
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   288
  \end{tabular}
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   289
  \medskip
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   290
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   291
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   292
    (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * )
26782
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
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 35613
diff changeset
   295
    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')'
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   296
    ;
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   297
    @{syntax_def simpmod}: ('add' | 'del' | 'only' | 'split' (() | 'add' | 'del') |
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   298
      'cong' (() | 'add' | 'del')) ':' @{syntax thmrefs}
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   299
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   300
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   301
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   302
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   303
  \item @{method simp} invokes the Simplifier on the first subgoal,
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   304
  after inserting chained facts as additional goal premises; further
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   305
  rule declarations may be included via @{text "(simp add: facts)"}.
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   306
  The proof method fails if the subgoal remains unchanged after
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   307
  simplification.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   308
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   309
  Note that the original goal premises and chained facts are subject
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   310
  to simplification themselves, while declarations via @{text
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   311
  "add"}/@{text "del"} merely follow the policies of the object-logic
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   312
  to extract rewrite rules from theorems, without further
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   313
  simplification.  This may lead to slightly different behavior in
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   314
  either case, which might be required precisely like that in some
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   315
  boundary situations to perform the intended simplification step!
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   316
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   317
  \medskip The @{text only} modifier first removes all other rewrite
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   318
  rules, looper tactics (including split rules), congruence rules, and
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   319
  then behaves like @{text add}.  Implicit solvers remain, which means
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   320
  that trivial rules like reflexivity or introduction of @{text
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   321
  "True"} are available to solve the simplified subgoals, but also
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   322
  non-trivial tools like linear arithmetic in HOL.  The latter may
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   323
  lead to some surprise of the meaning of ``only'' in Isabelle/HOL
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   324
  compared to English!
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   325
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   326
  \medskip The @{text split} modifiers add or delete rules for the
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   327
  Splitter (see also \secref{sec:simp-strategies} on the looper).
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   328
  This works only if the Simplifier method has been properly setup to
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   329
  include the Splitter (all major object logics such HOL, HOLCF, FOL,
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   330
  ZF do this already).
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   331
50065
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   332
  There is also a separate @{method_ref split} method available for
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   333
  single-step case splitting.  The effect of repeatedly applying
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   334
  @{text "(split thms)"} can be imitated by ``@{text "(simp only:
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   335
  split: thms)"}''.
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   336
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   337
  \medskip The @{text cong} modifiers add or delete Simplifier
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   338
  congruence rules (see also \secref{sec:simp-rules}); the default is
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   339
  to add.
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   340
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   341
  \item @{method simp_all} is similar to @{method simp}, but acts on
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   342
  all goals, working backwards from the last to the first one as usual
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   343
  in Isabelle.\footnote{The order is irrelevant for goals without
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   344
  schematic variables, so simplification might actually be performed
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   345
  in parallel here.}
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   346
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   347
  Chained facts are inserted into all subgoals, before the
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   348
  simplification process starts.  Further rule declarations are the
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   349
  same as for @{method simp}.
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   350
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   351
  The proof method fails if all subgoals remain unchanged after
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   352
  simplification.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   353
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   354
  \item @{attribute simp_depth_limit} limits the number of recursive
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   355
  invocations of the Simplifier during conditional rewriting.
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   356
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   357
  \end{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   358
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   359
  By default the Simplifier methods above take local assumptions fully
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   360
  into account, using equational assumptions in the subsequent
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   361
  normalization process, or simplifying assumptions themselves.
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   362
  Further options allow to fine-tune the behavior of the Simplifier
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   363
  in this respect, corresponding to a variety of ML tactics as
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   364
  follows.\footnote{Unlike the corresponding Isar proof methods, the
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   365
  ML tactics do not insist in changing the goal state.}
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   366
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   367
  \begin{center}
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   368
  \small
59782
wenzelm
parents: 59780
diff changeset
   369
  \begin{tabular}{|l|l|p{0.3\textwidth}|}
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   370
  \hline
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   371
  Isar method & ML tactic & behavior \\\hline
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   372
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   373
  @{text "(simp (no_asm))"} & @{ML simp_tac} & assumptions are ignored
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   374
  completely \\\hline
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   375
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   376
  @{text "(simp (no_asm_simp))"} & @{ML asm_simp_tac} & assumptions
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   377
  are used in the simplification of the conclusion but are not
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   378
  themselves simplified \\\hline
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   379
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   380
  @{text "(simp (no_asm_use))"} & @{ML full_simp_tac} & assumptions
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   381
  are simplified but are not used in the simplification of each other
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   382
  or the conclusion \\\hline
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   383
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   384
  @{text "(simp)"} & @{ML asm_full_simp_tac} & assumptions are used in
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   385
  the simplification of the conclusion and to simplify other
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   386
  assumptions \\\hline
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   387
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   388
  @{text "(simp (asm_lr))"} & @{ML asm_lr_simp_tac} & compatibility
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   389
  mode: an assumption is only used for simplifying assumptions which
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   390
  are to the right of it \\\hline
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   391
59782
wenzelm
parents: 59780
diff changeset
   392
  \end{tabular}
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   393
  \end{center}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   394
\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   395
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   396
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   397
subsubsection \<open>Examples\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   398
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   399
text \<open>We consider basic algebraic simplifications in Isabelle/HOL.
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   400
  The rather trivial goal @{prop "0 + (x + 0) = x + 0 + 0"} looks like
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   401
  a good candidate to be solved by a single call of @{method simp}:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   402
\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   403
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   404
lemma "0 + (x + 0) = x + 0 + 0" apply simp? oops
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   405
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   406
text \<open>The above attempt \emph{fails}, because @{term "0"} and @{term
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   407
  "op +"} in the HOL library are declared as generic type class
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   408
  operations, without stating any algebraic laws yet.  More specific
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   409
  types are required to get access to certain standard simplifications
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   410
  of the theory context, e.g.\ like this:\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   411
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   412
lemma fixes x :: nat shows "0 + (x + 0) = x + 0 + 0" by simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   413
lemma fixes x :: int shows "0 + (x + 0) = x + 0 + 0" by simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   414
lemma fixes x :: "'a :: monoid_add" shows "0 + (x + 0) = x + 0 + 0" by simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   415
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   416
text \<open>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   417
  \medskip In many cases, assumptions of a subgoal are also needed in
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   418
  the simplification process.  For example:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   419
\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   420
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   421
lemma fixes x :: nat shows "x = 0 \<Longrightarrow> x + x = 0" by simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   422
lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" apply simp oops
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   423
lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" using assms by simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   424
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   425
text \<open>As seen above, local assumptions that shall contribute to
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   426
  simplification need to be part of the subgoal already, or indicated
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   427
  explicitly for use by the subsequent method invocation.  Both too
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   428
  little or too much information can make simplification fail, for
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   429
  different reasons.
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   430
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   431
  In the next example the malicious assumption @{prop "\<And>x::nat. f x =
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   432
  g (f (g x))"} does not contribute to solve the problem, but makes
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   433
  the default @{method simp} method loop: the rewrite rule @{text "f
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   434
  ?x \<equiv> g (f (g ?x))"} extracted from the assumption does not
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   435
  terminate.  The Simplifier notices certain simple forms of
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   436
  nontermination, but not this one.  The problem can be solved
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   437
  nonetheless, by ignoring assumptions via special options as
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   438
  explained before:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   439
\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   440
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   441
lemma "(\<And>x::nat. f x = g (f (g x))) \<Longrightarrow> f 0 = f 0 + 0"
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   442
  by (simp (no_asm))
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   443
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   444
text \<open>The latter form is typical for long unstructured proof
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   445
  scripts, where the control over the goal content is limited.  In
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   446
  structured proofs it is usually better to avoid pushing too many
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   447
  facts into the goal state in the first place.  Assumptions in the
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   448
  Isar proof context do not intrude the reasoning if not used
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   449
  explicitly.  This is illustrated for a toplevel statement and a
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   450
  local proof body as follows:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   451
\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   452
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   453
lemma
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   454
  assumes "\<And>x::nat. f x = g (f (g x))"
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   455
  shows "f 0 = f 0 + 0" by simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   456
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   457
notepad
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   458
begin
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   459
  assume "\<And>x::nat. f x = g (f (g x))"
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   460
  have "f 0 = f 0 + 0" by simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   461
end
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   462
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   463
text \<open>\medskip Because assumptions may simplify each other, there
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   464
  can be very subtle cases of nontermination. For example, the regular
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   465
  @{method simp} method applied to @{prop "P (f x) \<Longrightarrow> y = x \<Longrightarrow> f x = f y
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   466
  \<Longrightarrow> Q"} gives rise to the infinite reduction sequence
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   467
  \[
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   468
  @{text "P (f x)"} \stackrel{@{text "f x \<equiv> f y"}}{\longmapsto}
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   469
  @{text "P (f y)"} \stackrel{@{text "y \<equiv> x"}}{\longmapsto}
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   470
  @{text "P (f x)"} \stackrel{@{text "f x \<equiv> f y"}}{\longmapsto} \cdots
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   471
  \]
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   472
  whereas applying the same to @{prop "y = x \<Longrightarrow> f x = f y \<Longrightarrow> P (f x) \<Longrightarrow>
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   473
  Q"} terminates (without solving the goal):
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   474
\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   475
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   476
lemma "y = x \<Longrightarrow> f x = f y \<Longrightarrow> P (f x) \<Longrightarrow> Q"
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   477
  apply simp
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   478
  oops
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   479
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   480
text \<open>See also \secref{sec:simp-trace} for options to enable
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   481
  Simplifier trace mode, which often helps to diagnose problems with
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   482
  rewrite systems.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   483
\<close>
50064
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   484
e08cc8b20564 refurbished Simplifier examples;
wenzelm
parents: 50063
diff changeset
   485
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   486
subsection \<open>Declaring rules \label{sec:simp-rules}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   487
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   488
text \<open>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   489
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   490
    @{attribute_def simp} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   491
    @{attribute_def split} & : & @{text attribute} \\
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   492
    @{attribute_def cong} & : & @{text attribute} \\
50077
wenzelm
parents: 50076
diff changeset
   493
    @{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   494
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   495
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   496
  @{rail \<open>
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   497
    (@@{attribute simp} | @@{attribute split} | @@{attribute cong})
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   498
      (() | 'add' | 'del')
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   499
    ;
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   500
    @@{command print_simpset} ('!'?)
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   501
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   502
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   503
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   504
50076
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   505
  \item @{attribute simp} declares rewrite rules, by adding or
50065
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   506
  deleting them from the simpset within the theory or proof context.
50076
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   507
  Rewrite rules are theorems expressing some form of equality, for
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   508
  example:
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   509
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   510
  @{text "Suc ?m + ?n = ?m + Suc ?n"} \\
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   511
  @{text "?P \<and> ?P \<longleftrightarrow> ?P"} \\
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   512
  @{text "?A \<union> ?B \<equiv> {x. x \<in> ?A \<or> x \<in> ?B}"}
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   513
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   514
  \smallskip
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   515
  Conditional rewrites such as @{text "?m < ?n \<Longrightarrow> ?m div ?n = 0"} are
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   516
  also permitted; the conditions can be arbitrary formulas.
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   517
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   518
  \medskip Internally, all rewrite rules are translated into Pure
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   519
  equalities, theorems with conclusion @{text "lhs \<equiv> rhs"}. The
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   520
  simpset contains a function for extracting equalities from arbitrary
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   521
  theorems, which is usually installed when the object-logic is
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   522
  configured initially. For example, @{text "\<not> ?x \<in> {}"} could be
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   523
  turned into @{text "?x \<in> {} \<equiv> False"}. Theorems that are declared as
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   524
  @{attribute simp} and local assumptions within a goal are treated
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   525
  uniformly in this respect.
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   526
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   527
  The Simplifier accepts the following formats for the @{text "lhs"}
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   528
  term:
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   529
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   530
  \begin{enumerate}
50065
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   531
50076
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   532
  \item First-order patterns, considering the sublanguage of
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   533
  application of constant operators to variable operands, without
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   534
  @{text "\<lambda>"}-abstractions or functional variables.
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   535
  For example:
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   536
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   537
  @{text "(?x + ?y) + ?z \<equiv> ?x + (?y + ?z)"} \\
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   538
  @{text "f (f ?x ?y) ?z \<equiv> f ?x (f ?y ?z)"}
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   539
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58310
diff changeset
   540
  \item Higher-order patterns in the sense of @{cite "nipkow-patterns"}.
50076
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   541
  These are terms in @{text "\<beta>"}-normal form (this will always be the
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   542
  case unless you have done something strange) where each occurrence
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   543
  of an unknown is of the form @{text "?F x\<^sub>1 \<dots> x\<^sub>n"}, where the
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   544
  @{text "x\<^sub>i"} are distinct bound variables.
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   545
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   546
  For example, @{text "(\<forall>x. ?P x \<and> ?Q x) \<equiv> (\<forall>x. ?P x) \<and> (\<forall>x. ?Q x)"}
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   547
  or its symmetric form, since the @{text "rhs"} is also a
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   548
  higher-order pattern.
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   549
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   550
  \item Physical first-order patterns over raw @{text "\<lambda>"}-term
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   551
  structure without @{text "\<alpha>\<beta>\<eta>"}-equality; abstractions and bound
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   552
  variables are treated like quasi-constant term material.
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   553
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   554
  For example, the rule @{text "?f ?x \<in> range ?f = True"} rewrites the
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   555
  term @{text "g a \<in> range g"} to @{text "True"}, but will fail to
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   556
  match @{text "g (h b) \<in> range (\<lambda>x. g (h x))"}. However, offending
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   557
  subterms (in our case @{text "?f ?x"}, which is not a pattern) can
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   558
  be replaced by adding new variables and conditions like this: @{text
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   559
  "?y = ?f ?x \<Longrightarrow> ?y \<in> range ?f = True"} is acceptable as a conditional
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   560
  rewrite rule of the second category since conditions can be
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   561
  arbitrary terms.
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   562
c5cc24781cbf updated explanation of rewrite rules;
wenzelm
parents: 50075
diff changeset
   563
  \end{enumerate}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   564
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   565
  \item @{attribute split} declares case split rules.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   566
45645
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   567
  \item @{attribute cong} declares congruence rules to the Simplifier
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   568
  context.
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   569
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   570
  Congruence rules are equalities of the form @{text [display]
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   571
  "\<dots> \<Longrightarrow> f ?x\<^sub>1 \<dots> ?x\<^sub>n = f ?y\<^sub>1 \<dots> ?y\<^sub>n"}
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   572
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   573
  This controls the simplification of the arguments of @{text f}.  For
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   574
  example, some arguments can be simplified under additional
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   575
  assumptions: @{text [display] "?P\<^sub>1 \<longleftrightarrow> ?Q\<^sub>1 \<Longrightarrow> (?Q\<^sub>1 \<Longrightarrow> ?P\<^sub>2 \<longleftrightarrow> ?Q\<^sub>2) \<Longrightarrow>
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   576
  (?P\<^sub>1 \<longrightarrow> ?P\<^sub>2) \<longleftrightarrow> (?Q\<^sub>1 \<longrightarrow> ?Q\<^sub>2)"}
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   577
56594
e3a06699a13f tuned spelling;
wenzelm
parents: 56451
diff changeset
   578
  Given this rule, the Simplifier assumes @{text "?Q\<^sub>1"} and extracts
45645
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   579
  rewrite rules from it when simplifying @{text "?P\<^sub>2"}.  Such local
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   580
  assumptions are effective for rewriting formulae such as @{text "x =
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   581
  0 \<longrightarrow> y + x = y"}.
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   582
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   583
  %FIXME
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   584
  %The local assumptions are also provided as theorems to the solver;
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   585
  %see \secref{sec:simp-solver} below.
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   586
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   587
  \medskip The following congruence rule for bounded quantifiers also
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   588
  supplies contextual information --- about the bound variable:
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   589
  @{text [display] "(?A = ?B) \<Longrightarrow> (\<And>x. x \<in> ?B \<Longrightarrow> ?P x \<longleftrightarrow> ?Q x) \<Longrightarrow>
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   590
    (\<forall>x \<in> ?A. ?P x) \<longleftrightarrow> (\<forall>x \<in> ?B. ?Q x)"}
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   591
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   592
  \medskip This congruence rule for conditional expressions can
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   593
  supply contextual information for simplifying the arms:
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   594
  @{text [display] "?p = ?q \<Longrightarrow> (?q \<Longrightarrow> ?a = ?c) \<Longrightarrow> (\<not> ?q \<Longrightarrow> ?b = ?d) \<Longrightarrow>
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   595
    (if ?p then ?a else ?b) = (if ?q then ?c else ?d)"}
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   596
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   597
  A congruence rule can also \emph{prevent} simplification of some
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   598
  arguments.  Here is an alternative congruence rule for conditional
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   599
  expressions that conforms to non-strict functional evaluation:
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   600
  @{text [display] "?p = ?q \<Longrightarrow> (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"}
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   601
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   602
  Only the first argument is simplified; the others remain unchanged.
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   603
  This can make simplification much faster, but may require an extra
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   604
  case split over the condition @{text "?q"} to prove the goal.
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   605
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   606
  \item @{command "print_simpset"} prints the collection of rules declared
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   607
  to the Simplifier, which is also known as ``simpset'' internally; the
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59905
diff changeset
   608
  ``@{text "!"}'' option indicates extra verbosity.
50077
wenzelm
parents: 50076
diff changeset
   609
wenzelm
parents: 50076
diff changeset
   610
  For historical reasons, simpsets may occur independently from the
wenzelm
parents: 50076
diff changeset
   611
  current context, but are conceptually dependent on it.  When the
wenzelm
parents: 50076
diff changeset
   612
  Simplifier is invoked via one of its main entry points in the Isar
wenzelm
parents: 50076
diff changeset
   613
  source language (as proof method \secref{sec:simp-meth} or rule
wenzelm
parents: 50076
diff changeset
   614
  attribute \secref{sec:simp-meth}), its simpset is derived from the
wenzelm
parents: 50076
diff changeset
   615
  current proof context, and carries a back-reference to that for
wenzelm
parents: 50076
diff changeset
   616
  other tools that might get invoked internally (e.g.\ simplification
wenzelm
parents: 50076
diff changeset
   617
  procedures \secref{sec:simproc}).  A mismatch of the context of the
wenzelm
parents: 50076
diff changeset
   618
  simpset and the context of the problem being simplified may lead to
wenzelm
parents: 50076
diff changeset
   619
  unexpected results.
wenzelm
parents: 50076
diff changeset
   620
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   621
  \end{description}
50065
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   622
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   623
  The implicit simpset of the theory context is propagated
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   624
  monotonically through the theory hierarchy: forming a new theory,
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   625
  the union of the simpsets of its imports are taken as starting
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   626
  point.  Also note that definitional packages like @{command
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   627
  "datatype"}, @{command "primrec"}, @{command "fun"} routinely
50065
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   628
  declare Simplifier rules to the target context, while plain
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   629
  @{command "definition"} is an exception in \emph{not} declaring
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   630
  anything.
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   631
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   632
  \medskip It is up the user to manipulate the current simpset further
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   633
  by explicitly adding or deleting theorems as simplification rules,
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   634
  or installing other tools via simplification procedures
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   635
  (\secref{sec:simproc}).  Good simpsets are hard to design.  Rules
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   636
  that obviously simplify, like @{text "?n + 0 \<equiv> ?n"} are good
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   637
  candidates for the implicit simpset, unless a special
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   638
  non-normalizing behavior of certain operations is intended.  More
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   639
  specific rules (such as distributive laws, which duplicate subterms)
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   640
  should be added only for specific proof steps.  Conversely,
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   641
  sometimes a rule needs to be deleted just for some part of a proof.
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   642
  The need of frequent additions or deletions may indicate a poorly
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   643
  designed simpset.
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   644
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   645
  \begin{warn}
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   646
  The union of simpsets from theory imports (as described above) is
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   647
  not always a good starting point for the new theory.  If some
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   648
  ancestors have deleted simplification rules because they are no
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   649
  longer wanted, while others have left those rules in, then the union
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   650
  will contain the unwanted rules, and thus have to be deleted again
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   651
  in the theory body.
7c01dc2dcb8c more on Simplifier rules, based on old material;
wenzelm
parents: 50064
diff changeset
   652
  \end{warn}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   653
\<close>
45645
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   654
4014bc2a09ff modernized section about congruence rules;
wenzelm
parents: 44911
diff changeset
   655
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   656
subsection \<open>Ordered rewriting with permutative rules\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   657
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   658
text \<open>A rewrite rule is \emph{permutative} if the left-hand side and
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   659
  right-hand side are the equal up to renaming of variables.  The most
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   660
  common permutative rule is commutativity: @{text "?x + ?y = ?y +
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   661
  ?x"}.  Other examples include @{text "(?x - ?y) - ?z = (?x - ?z) -
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   662
  ?y"} in arithmetic and @{text "insert ?x (insert ?y ?A) = insert ?y
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   663
  (insert ?x ?A)"} for sets.  Such rules are common enough to merit
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   664
  special attention.
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   665
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   666
  Because ordinary rewriting loops given such rules, the Simplifier
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   667
  employs a special strategy, called \emph{ordered rewriting}.
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   668
  Permutative rules are detected and only applied if the rewriting
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   669
  step decreases the redex wrt.\ a given term ordering.  For example,
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   670
  commutativity rewrites @{text "b + a"} to @{text "a + b"}, but then
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   671
  stops, because the redex cannot be decreased further in the sense of
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   672
  the term ordering.
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   673
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   674
  The default is lexicographic ordering of term structure, but this
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   675
  could be also changed locally for special applications via
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   676
  @{index_ML Simplifier.set_termless} in Isabelle/ML.
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   677
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   678
  \medskip Permutative rewrite rules are declared to the Simplifier
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   679
  just like other rewrite rules.  Their special status is recognized
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   680
  automatically, and their application is guarded by the term ordering
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   681
  accordingly.\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   682
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   683
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   684
subsubsection \<open>Rewriting with AC operators\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   685
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   686
text \<open>Ordered rewriting is particularly effective in the case of
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   687
  associative-commutative operators.  (Associativity by itself is not
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   688
  permutative.)  When dealing with an AC-operator @{text "f"}, keep
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   689
  the following points in mind:
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   690
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   691
  \begin{itemize}
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   692
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   693
  \item The associative law must always be oriented from left to
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   694
  right, namely @{text "f (f x y) z = f x (f y z)"}.  The opposite
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   695
  orientation, if used with commutativity, leads to looping in
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   696
  conjunction with the standard term order.
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   697
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   698
  \item To complete your set of rewrite rules, you must add not just
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   699
  associativity (A) and commutativity (C) but also a derived rule
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   700
  \emph{left-commutativity} (LC): @{text "f x (f y z) = f y (f x z)"}.
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   701
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   702
  \end{itemize}
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   703
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   704
  Ordered rewriting with the combination of A, C, and LC sorts a term
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   705
  lexicographically --- the rewriting engine imitates bubble-sort.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   706
\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   707
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
   708
experiment
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   709
  fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infix "\<bullet>" 60)
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   710
  assumes assoc: "(x \<bullet> y) \<bullet> z = x \<bullet> (y \<bullet> z)"
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   711
  assumes commute: "x \<bullet> y = y \<bullet> x"
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   712
begin
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   713
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   714
lemma left_commute: "x \<bullet> (y \<bullet> z) = y \<bullet> (x \<bullet> z)"
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   715
proof -
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   716
  have "(x \<bullet> y) \<bullet> z = (y \<bullet> x) \<bullet> z" by (simp only: commute)
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   717
  then show ?thesis by (simp only: assoc)
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   718
qed
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   719
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   720
lemmas AC_rules = assoc commute left_commute
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   721
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   722
text \<open>Thus the Simplifier is able to establish equalities with
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   723
  arbitrary permutations of subterms, by normalizing to a common
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   724
  standard form.  For example:\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   725
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   726
lemma "(b \<bullet> c) \<bullet> a = xxx"
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   727
  apply (simp only: AC_rules)
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   728
  txt \<open>@{subgoals}\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   729
  oops
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   730
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   731
lemma "(b \<bullet> c) \<bullet> a = a \<bullet> (b \<bullet> c)" by (simp only: AC_rules)
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   732
lemma "(b \<bullet> c) \<bullet> a = c \<bullet> (b \<bullet> a)" by (simp only: AC_rules)
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   733
lemma "(b \<bullet> c) \<bullet> a = (c \<bullet> b) \<bullet> a" by (simp only: AC_rules)
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   734
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   735
end
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   736
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   737
text \<open>Martin and Nipkow @{cite "martin-nipkow"} discuss the theory and
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   738
  give many examples; other algebraic structures are amenable to
56594
e3a06699a13f tuned spelling;
wenzelm
parents: 56451
diff changeset
   739
  ordered rewriting, such as Boolean rings.  The Boyer-Moore theorem
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58310
diff changeset
   740
  prover @{cite bm88book} also employs ordered rewriting.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   741
\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   742
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   743
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   744
subsubsection \<open>Re-orienting equalities\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   745
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   746
text \<open>Another application of ordered rewriting uses the derived rule
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   747
  @{thm [source] eq_commute}: @{thm [source = false] eq_commute} to
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   748
  reverse equations.
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   749
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   750
  This is occasionally useful to re-orient local assumptions according
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   751
  to the term ordering, when other built-in mechanisms of
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   752
  reorientation and mutual simplification fail to apply.\<close>
50080
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   753
200f749c96db updated section on ordered rewriting;
wenzelm
parents: 50079
diff changeset
   754
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   755
subsection \<open>Simplifier tracing and debugging \label{sec:simp-trace}\<close>
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   756
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   757
text \<open>
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   758
  \begin{tabular}{rcll}
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   759
    @{attribute_def simp_trace} & : & @{text attribute} & default @{text false} \\
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   760
    @{attribute_def simp_trace_depth_limit} & : & @{text attribute} & default @{text 1} \\
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   761
    @{attribute_def simp_debug} & : & @{text attribute} & default @{text false} \\
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   762
    @{attribute_def simp_trace_new} & : & @{text attribute} \\
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   763
    @{attribute_def simp_break} & : & @{text attribute} \\
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   764
  \end{tabular}
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   765
  \medskip
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   766
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   767
  @{rail \<open>
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   768
    @@{attribute simp_trace_new} ('interactive')? \<newline>
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   769
      ('mode' '=' ('full' | 'normal'))? \<newline>
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   770
      ('depth' '=' @{syntax nat})?
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   771
    ;
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   772
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   773
    @@{attribute simp_break} (@{syntax term}*)
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   774
  \<close>}
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   775
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   776
  These attributes and configurations options control various aspects of
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   777
  Simplifier tracing and debugging.
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   778
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   779
  \begin{description}
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   780
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   781
  \item @{attribute simp_trace} makes the Simplifier output internal
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   782
  operations.  This includes rewrite steps, but also bookkeeping like
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   783
  modifications of the simpset.
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   784
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   785
  \item @{attribute simp_trace_depth_limit} limits the effect of
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   786
  @{attribute simp_trace} to the given depth of recursive Simplifier
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   787
  invocations (when solving conditions of rewrite rules).
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   788
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   789
  \item @{attribute simp_debug} makes the Simplifier output some extra
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   790
  information about internal operations.  This includes any attempted
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   791
  invocation of simplification procedures.
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   792
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   793
  \item @{attribute simp_trace_new} controls Simplifier tracing within
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58310
diff changeset
   794
  Isabelle/PIDE applications, notably Isabelle/jEdit @{cite "isabelle-jedit"}.
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   795
  This provides a hierarchical representation of the rewriting steps
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   796
  performed by the Simplifier.
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   797
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   798
  Users can configure the behaviour by specifying breakpoints, verbosity and
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   799
  enabling or disabling the interactive mode. In normal verbosity (the
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   800
  default), only rule applications matching a breakpoint will be shown to
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   801
  the user. In full verbosity, all rule applications will be logged.
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   802
  Interactive mode interrupts the normal flow of the Simplifier and defers
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   803
  the decision how to continue to the user via some GUI dialog.
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   804
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   805
  \item @{attribute simp_break} declares term or theorem breakpoints for
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   806
  @{attribute simp_trace_new} as described above. Term breakpoints are
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   807
  patterns which are checked for matches on the redex of a rule application.
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   808
  Theorem breakpoints trigger when the corresponding theorem is applied in a
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   809
  rewrite step. For example:
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   810
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   811
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   812
\<close>
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   813
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
   814
(*<*)experiment begin(*>*)
57591
8c095aef6769 clarified "simp_trace_new" and corresponding isar-ref section;
wenzelm
parents: 57590
diff changeset
   815
declare conjI [simp_break]
57590
06cb5375e189 more on "Simplifier trace" (by Lars Hupel);
wenzelm
parents: 56594
diff changeset
   816
declare [[simp_break "?x \<and> ?y"]]
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
   817
(*<*)end(*>*)
57590
06cb5375e189 more on "Simplifier trace" (by Lars Hupel);
wenzelm
parents: 56594
diff changeset
   818
50063
7e491da6bcbd more on the Simplifier, based on old material;
wenzelm
parents: 48985
diff changeset
   819
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   820
subsection \<open>Simplification procedures \label{sec:simproc}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   821
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   822
text \<open>Simplification procedures are ML functions that produce proven
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   823
  rewrite rules on demand.  They are associated with higher-order
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   824
  patterns that approximate the left-hand sides of equations.  The
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   825
  Simplifier first matches the current redex against one of the LHS
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   826
  patterns; if this succeeds, the corresponding ML function is
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   827
  invoked, passing the Simplifier context and redex term.  Thus rules
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   828
  may be specifically fashioned for particular situations, resulting
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   829
  in a more powerful mechanism than term rewriting by a fixed set of
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   830
  rules.
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   831
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   832
  Any successful result needs to be a (possibly conditional) rewrite
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   833
  rule @{text "t \<equiv> u"} that is applicable to the current redex.  The
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   834
  rule will be applied just as any ordinary rewrite rule.  It is
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   835
  expected to be already in \emph{internal form}, bypassing the
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   836
  automatic preprocessing of object-level equivalences.
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   837
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   838
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
   839
    @{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
   840
    simproc & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   841
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   842
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   843
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   844
    @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '='
55029
61a6bf7d4b02 clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents: 52060
diff changeset
   845
      @{syntax text} \<newline> (@'identifier' (@{syntax nameref}+))?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   846
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   847
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
   848
    @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+)
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   849
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   850
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   851
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   852
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   853
  \item @{command "simproc_setup"} defines a named simplification
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   854
  procedure that is invoked by the Simplifier whenever any of the
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   855
  given term patterns match the current redex.  The implementation,
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   856
  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
   857
  "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   858
  cterm} represents the current redex @{text r} and the result is
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   859
  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
   860
  generalized version), or @{ML NONE} to indicate failure.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   861
  @{ML_type simpset} argument holds the full context of the current
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   862
  Simplifier invocation, including the actual Isar proof context.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   863
  @{ML_type morphism} informs about the difference of the original
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   864
  compilation context wrt.\ the one of the actual application later
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   865
  on.  The optional @{keyword "identifier"} specifies theorems that
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   866
  represent the logical content of the abstract theory of this
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   867
  simproc.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   868
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   869
  Morphisms and identifiers are only relevant for simprocs that are
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   870
  defined within a local target context, e.g.\ in a locale.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   871
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   872
  \item @{text "simproc add: name"} and @{text "simproc del: name"}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   873
  add or delete named simprocs to the current Simplifier context.  The
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   874
  default is to add a simproc.  Note that @{command "simproc_setup"}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   875
  already adds the new simproc to the subsequent context.
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   876
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
   877
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   878
\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   879
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
   880
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   881
subsubsection \<open>Example\<close>
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   882
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   883
text \<open>The following simplification procedure for @{thm
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   884
  [source=false, show_types] unit_eq} in HOL performs fine-grained
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   885
  control over rule application, beyond higher-order pattern matching.
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   886
  Declaring @{thm unit_eq} as @{attribute simp} directly would make
56594
e3a06699a13f tuned spelling;
wenzelm
parents: 56451
diff changeset
   887
  the Simplifier loop!  Note that a version of this simplification
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   888
  procedure is already active in Isabelle/HOL.\<close>
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   889
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
   890
(*<*)experiment begin(*>*)
59782
wenzelm
parents: 59780
diff changeset
   891
simproc_setup unit ("x::unit") =
wenzelm
parents: 59780
diff changeset
   892
  \<open>fn _ => fn _ => fn ct =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   893
    if HOLogic.is_unit (Thm.term_of ct) then NONE
59782
wenzelm
parents: 59780
diff changeset
   894
    else SOME (mk_meta_eq @{thm unit_eq})\<close>
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
   895
(*<*)end(*>*)
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   896
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   897
text \<open>Since the Simplifier applies simplification procedures
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   898
  frequently, it is important to make the failure check in ML
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   899
  reasonably fast.\<close>
42925
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   900
c6c4f997ad87 updated and re-unified material on simprocs;
wenzelm
parents: 42705
diff changeset
   901
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   902
subsection \<open>Configurable Simplifier strategies \label{sec:simp-strategies}\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   903
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   904
text \<open>The core term-rewriting engine of the Simplifier is normally
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   905
  used in combination with some add-on components that modify the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   906
  strategy and allow to integrate other non-Simplifier proof tools.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   907
  These may be reconfigured in ML as explained below.  Even if the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   908
  default strategies of object-logics like Isabelle/HOL are used
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   909
  unchanged, it helps to understand how the standard Simplifier
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   910
  strategies work.\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   911
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   912
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   913
subsubsection \<open>The subgoaler\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   914
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   915
text \<open>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   916
  \begin{mldecls}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   917
  @{index_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   918
  Proof.context -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   919
  @{index_ML Simplifier.prems_of: "Proof.context -> thm list"} \\
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   920
  \end{mldecls}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   921
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   922
  The subgoaler is the tactic used to solve subgoals arising out of
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   923
  conditional rewrite rules or congruence rules.  The default should
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   924
  be simplification itself.  In rare situations, this strategy may
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   925
  need to be changed.  For example, if the premise of a conditional
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   926
  rule is an instance of its conclusion, as in @{text "Suc ?m < ?n \<Longrightarrow>
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   927
  ?m < ?n"}, the default strategy could loop.  % FIXME !??
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   928
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   929
  \begin{description}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   930
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   931
  \item @{ML Simplifier.set_subgoaler}~@{text "tac ctxt"} sets the
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   932
  subgoaler of the context to @{text "tac"}.  The tactic will
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   933
  be applied to the context of the running Simplifier instance.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   934
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   935
  \item @{ML Simplifier.prems_of}~@{text "ctxt"} retrieves the current
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   936
  set of premises from the context.  This may be non-empty only if
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   937
  the Simplifier has been told to utilize local assumptions in the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   938
  first place (cf.\ the options in \secref{sec:simp-meth}).
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   939
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   940
  \end{description}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   941
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   942
  As an example, consider the following alternative subgoaler:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   943
\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   944
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
   945
ML_val \<open>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   946
  fun subgoaler_tac ctxt =
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58618
diff changeset
   947
    assume_tac ctxt ORELSE'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   948
    resolve_tac ctxt (Simplifier.prems_of ctxt) ORELSE'
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   949
    asm_simp_tac ctxt
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   950
\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   951
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   952
text \<open>This tactic first tries to solve the subgoal by assumption or
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   953
  by resolving with with one of the premises, calling simplification
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   954
  only if that fails.\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   955
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   956
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   957
subsubsection \<open>The solver\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   958
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   959
text \<open>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   960
  \begin{mldecls}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   961
  @{index_ML_type solver} \\
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   962
  @{index_ML Simplifier.mk_solver: "string ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   963
  (Proof.context -> int -> tactic) -> solver"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   964
  @{index_ML_op setSolver: "Proof.context * solver -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   965
  @{index_ML_op addSolver: "Proof.context * solver -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   966
  @{index_ML_op setSSolver: "Proof.context * solver -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   967
  @{index_ML_op addSSolver: "Proof.context * solver -> Proof.context"} \\
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   968
  \end{mldecls}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   969
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   970
  A solver is a tactic that attempts to solve a subgoal after
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   971
  simplification.  Its core functionality is to prove trivial subgoals
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   972
  such as @{prop "True"} and @{text "t = t"}, but object-logics might
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   973
  be more ambitious.  For example, Isabelle/HOL performs a restricted
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   974
  version of linear arithmetic here.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   975
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   976
  Solvers are packaged up in abstract type @{ML_type solver}, with
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   977
  @{ML Simplifier.mk_solver} as the only operation to create a solver.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   978
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   979
  \medskip Rewriting does not instantiate unknowns.  For example,
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   980
  rewriting alone cannot prove @{text "a \<in> ?A"} since this requires
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   981
  instantiating @{text "?A"}.  The solver, however, is an arbitrary
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   982
  tactic and may instantiate unknowns as it pleases.  This is the only
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   983
  way the Simplifier can handle a conditional rewrite rule whose
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   984
  condition contains extra variables.  When a simplification tactic is
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   985
  to be combined with other provers, especially with the Classical
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   986
  Reasoner, it is important whether it can be considered safe or not.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   987
  For this reason a simpset contains two solvers: safe and unsafe.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   988
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   989
  The standard simplification strategy solely uses the unsafe solver,
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   990
  which is appropriate in most cases.  For special applications where
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   991
  the simplification process is not allowed to instantiate unknowns
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   992
  within the goal, simplification starts with the safe solver, but may
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   993
  still apply the ordinary unsafe one in nested simplifications for
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   994
  conditional rules or congruences. Note that in this way the overall
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   995
  tactic is not totally safe: it may instantiate unknowns that appear
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   996
  also in other subgoals.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   997
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   998
  \begin{description}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
   999
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1000
  \item @{ML Simplifier.mk_solver}~@{text "name tac"} turns @{text
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1001
  "tac"} into a solver; the @{text "name"} is only attached as a
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1002
  comment and has no further significance.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1003
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1004
  \item @{text "ctxt setSSolver solver"} installs @{text "solver"} as
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1005
  the safe solver of @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1006
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1007
  \item @{text "ctxt addSSolver solver"} adds @{text "solver"} as an
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1008
  additional safe solver; it will be tried after the solvers which had
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1009
  already been present in @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1010
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1011
  \item @{text "ctxt setSolver solver"} installs @{text "solver"} as the
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1012
  unsafe solver of @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1013
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1014
  \item @{text "ctxt addSolver solver"} adds @{text "solver"} as an
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1015
  additional unsafe solver; it will be tried after the solvers which
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1016
  had already been present in @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1017
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1018
  \end{description}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1019
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1020
  \medskip The solver tactic is invoked with the context of the
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1021
  running Simplifier.  Further operations
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1022
  may be used to retrieve relevant information, such as the list of
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1023
  local Simplifier premises via @{ML Simplifier.prems_of} --- this
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1024
  list may be non-empty only if the Simplifier runs in a mode that
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1025
  utilizes local assumptions (see also \secref{sec:simp-meth}).  The
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1026
  solver is also presented the full goal including its assumptions in
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1027
  any case.  Thus it can use these (e.g.\ by calling @{ML
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1028
  assume_tac}), even if the Simplifier proper happens to ignore local
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1029
  premises at the moment.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1030
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1031
  \medskip As explained before, the subgoaler is also used to solve
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1032
  the premises of congruence rules.  These are usually of the form
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1033
  @{text "s = ?x"}, where @{text "s"} needs to be simplified and
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1034
  @{text "?x"} needs to be instantiated with the result.  Typically,
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1035
  the subgoaler will invoke the Simplifier at some point, which will
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1036
  eventually call the solver.  For this reason, solver tactics must be
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1037
  prepared to solve goals of the form @{text "t = ?x"}, usually by
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1038
  reflexivity.  In particular, reflexivity should be tried before any
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1039
  of the fancy automated proof tools.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1040
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1041
  It may even happen that due to simplification the subgoal is no
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1042
  longer an equality.  For example, @{text "False \<longleftrightarrow> ?Q"} could be
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1043
  rewritten to @{text "\<not> ?Q"}.  To cover this case, the solver could
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1044
  try resolving with the theorem @{text "\<not> False"} of the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1045
  object-logic.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1046
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1047
  \medskip
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1048
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1049
  \begin{warn}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1050
  If a premise of a congruence rule cannot be proved, then the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1051
  congruence is ignored.  This should only happen if the rule is
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1052
  \emph{conditional} --- that is, contains premises not of the form
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1053
  @{text "t = ?x"}.  Otherwise it indicates that some congruence rule,
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1054
  or possibly the subgoaler or solver, is faulty.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1055
  \end{warn}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1056
\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1057
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1058
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1059
subsubsection \<open>The looper\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1060
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1061
text \<open>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1062
  \begin{mldecls}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1063
  @{index_ML_op setloop: "Proof.context *
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1064
  (Proof.context -> int -> tactic) -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1065
  @{index_ML_op addloop: "Proof.context *
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1066
  (string * (Proof.context -> int -> tactic))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1067
  -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1068
  @{index_ML_op delloop: "Proof.context * string -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1069
  @{index_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1070
  @{index_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1071
  \end{mldecls}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1072
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1073
  The looper is a list of tactics that are applied after
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1074
  simplification, in case the solver failed to solve the simplified
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1075
  goal.  If the looper succeeds, the simplification process is started
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1076
  all over again.  Each of the subgoals generated by the looper is
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1077
  attacked in turn, in reverse order.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1078
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1079
  A typical looper is \emph{case splitting}: the expansion of a
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1080
  conditional.  Another possibility is to apply an elimination rule on
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1081
  the assumptions.  More adventurous loopers could start an induction.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1082
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1083
  \begin{description}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1084
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1085
  \item @{text "ctxt setloop tac"} installs @{text "tac"} as the only
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
  1086
  looper tactic of @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1087
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1088
  \item @{text "ctxt addloop (name, tac)"} adds @{text "tac"} as an
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1089
  additional looper tactic with name @{text "name"}, which is
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1090
  significant for managing the collection of loopers.  The tactic will
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1091
  be tried after the looper tactics that had already been present in
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
  1092
  @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1093
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1094
  \item @{text "ctxt delloop name"} deletes the looper tactic that was
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1095
  associated with @{text "name"} from @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1096
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1097
  \item @{ML Splitter.add_split}~@{text "thm ctxt"} adds split tactics
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1098
  for @{text "thm"} as additional looper tactics of @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1099
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1100
  \item @{ML Splitter.del_split}~@{text "thm ctxt"} deletes the split
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1101
  tactic corresponding to @{text thm} from the looper tactics of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1102
  @{text "ctxt"}.
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1103
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1104
  \end{description}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1105
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1106
  The splitter replaces applications of a given function; the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1107
  right-hand side of the replacement can be anything.  For example,
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1108
  here is a splitting rule for conditional expressions:
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1109
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1110
  @{text [display] "?P (if ?Q ?x ?y) \<longleftrightarrow> (?Q \<longrightarrow> ?P ?x) \<and> (\<not> ?Q \<longrightarrow> ?P ?y)"}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1111
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1112
  Another example is the elimination operator for Cartesian products
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1113
  (which happens to be called @{text split} in Isabelle/HOL:
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1114
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1115
  @{text [display] "?P (split ?f ?p) \<longleftrightarrow> (\<forall>a b. ?p = (a, b) \<longrightarrow> ?P (f a b))"}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1116
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1117
  For technical reasons, there is a distinction between case splitting
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1118
  in the conclusion and in the premises of a subgoal.  The former is
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1119
  done by @{ML Splitter.split_tac} with rules like @{thm [source]
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1120
  split_if} or @{thm [source] option.split}, which do not split the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1121
  subgoal, while the latter is done by @{ML Splitter.split_asm_tac}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1122
  with rules like @{thm [source] split_if_asm} or @{thm [source]
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1123
  option.split_asm}, which split the subgoal.  The function @{ML
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1124
  Splitter.add_split} automatically takes care of which tactic to
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1125
  call, analyzing the form of the rules given as argument; it is the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1126
  same operation behind @{text "split"} attribute or method modifier
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1127
  syntax in the Isar source language.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1128
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1129
  Case splits should be allowed only when necessary; they are
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1130
  expensive and hard to control.  Case-splitting on if-expressions in
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1131
  the conclusion is usually beneficial, so it is enabled by default in
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1132
  Isabelle/HOL and Isabelle/FOL/ZF.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1133
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1134
  \begin{warn}
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1135
  With @{ML Splitter.split_asm_tac} as looper component, the
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1136
  Simplifier may split subgoals!  This might cause unexpected problems
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1137
  in tactic expressions that silently assume 0 or 1 subgoals after
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1138
  simplification.
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1139
  \end{warn}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1140
\<close>
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1141
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1142
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1143
subsection \<open>Forward simplification \label{sec:simp-forward}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1144
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1145
text \<open>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1146
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1147
    @{attribute_def simplified} & : & @{text attribute} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1148
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1149
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1150
  @{rail \<open>
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1151
    @@{attribute simplified} opt? @{syntax thmrefs}?
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1152
    ;
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1153
40255
9ffbc25e1606 eliminated obsolete \_ escapes in rail environments;
wenzelm
parents: 35613
diff changeset
  1154
    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')'
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1155
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1156
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1157
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1158
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1159
  \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
  1160
  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
  1161
  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
  1162
  The result is fully simplified by default, including assumptions and
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1163
  conclusion; the options @{text no_asm} etc.\ tune the Simplifier in
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1164
  the same way as the for the @{text simp} method.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1165
56594
e3a06699a13f tuned spelling;
wenzelm
parents: 56451
diff changeset
  1166
  Note that forward simplification restricts the Simplifier to its
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1167
  most basic operation of term rewriting; solver and looper tactics
50079
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1168
  (\secref{sec:simp-strategies}) are \emph{not} involved here.  The
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1169
  @{attribute simplified} attribute should be only rarely required
5c36db9db335 updated subgoaler/solver/looper;
wenzelm
parents: 50077
diff changeset
  1170
  under normal circumstances.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1171
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1172
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1173
\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1174
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1175
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1176
section \<open>The Classical Reasoner \label{sec:classical}\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1177
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1178
subsection \<open>Basic concepts\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1179
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1180
text \<open>Although Isabelle is generic, many users will be working in
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1181
  some extension of classical first-order logic.  Isabelle/ZF is built
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1182
  upon theory FOL, while Isabelle/HOL conceptually contains
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1183
  first-order logic as a fragment.  Theorem-proving in predicate logic
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1184
  is undecidable, but many automated strategies have been developed to
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1185
  assist in this task.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1186
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1187
  Isabelle's classical reasoner is a generic package that accepts
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1188
  certain information about a logic and delivers a suite of automatic
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1189
  proof tools, based on rules that are classified and declared in the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1190
  context.  These proof procedures are slow and simplistic compared
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1191
  with high-end automated theorem provers, but they can save
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1192
  considerable time and effort in practice.  They can prove theorems
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58310
diff changeset
  1193
  such as Pelletier's @{cite pelletier86} problems 40 and 41 in a few
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1194
  milliseconds (including full proof reconstruction):\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1195
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1196
lemma "(\<exists>y. \<forall>x. F x y \<longleftrightarrow> F x x) \<longrightarrow> \<not> (\<forall>x. \<exists>y. \<forall>z. F z y \<longleftrightarrow> \<not> F z x)"
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1197
  by blast
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1198
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1199
lemma "(\<forall>z. \<exists>y. \<forall>x. f x y \<longleftrightarrow> f x z \<and> \<not> f x x) \<longrightarrow> \<not> (\<exists>z. \<forall>x. f x z)"
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1200
  by blast
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1201
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1202
text \<open>The proof tools are generic.  They are not restricted to
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1203
  first-order logic, and have been heavily used in the development of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1204
  the Isabelle/HOL library and applications.  The tactics can be
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1205
  traced, and their components can be called directly; in this manner,
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1206
  any proof can be viewed interactively.\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1207
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1208
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1209
subsubsection \<open>The sequent calculus\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1210
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1211
text \<open>Isabelle supports natural deduction, which is easy to use for
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1212
  interactive proof.  But natural deduction does not easily lend
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1213
  itself to automation, and has a bias towards intuitionism.  For
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1214
  certain proofs in classical logic, it can not be called natural.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1215
  The \emph{sequent calculus}, a generalization of natural deduction,
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1216
  is easier to automate.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1217
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1218
  A \textbf{sequent} has the form @{text "\<Gamma> \<turnstile> \<Delta>"}, where @{text "\<Gamma>"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1219
  and @{text "\<Delta>"} are sets of formulae.\footnote{For first-order
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1220
  logic, sequents can equivalently be made from lists or multisets of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1221
  formulae.} The sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} is
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1222
  \textbf{valid} if @{text "P\<^sub>1 \<and> \<dots> \<and> P\<^sub>m"} implies @{text "Q\<^sub>1 \<or> \<dots> \<or>
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1223
  Q\<^sub>n"}.  Thus @{text "P\<^sub>1, \<dots>, P\<^sub>m"} represent assumptions, each of which
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1224
  is true, while @{text "Q\<^sub>1, \<dots>, Q\<^sub>n"} represent alternative goals.  A
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1225
  sequent is \textbf{basic} if its left and right sides have a common
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1226
  formula, as in @{text "P, Q \<turnstile> Q, R"}; basic sequents are trivially
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1227
  valid.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1228
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1229
  Sequent rules are classified as \textbf{right} or \textbf{left},
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1230
  indicating which side of the @{text "\<turnstile>"} symbol they operate on.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1231
  Rules that operate on the right side are analogous to natural
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1232
  deduction's introduction rules, and left rules are analogous to
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1233
  elimination rules.  The sequent calculus analogue of @{text "(\<longrightarrow>I)"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1234
  is the rule
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1235
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1236
  \infer[@{text "(\<longrightarrow>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<longrightarrow> Q"}}{@{text "P, \<Gamma> \<turnstile> \<Delta>, Q"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1237
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1238
  Applying the rule backwards, this breaks down some implication on
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1239
  the right side of a sequent; @{text "\<Gamma>"} and @{text "\<Delta>"} stand for
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1240
  the sets of formulae that are unaffected by the inference.  The
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1241
  analogue of the pair @{text "(\<or>I1)"} and @{text "(\<or>I2)"} is the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1242
  single rule
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1243
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1244
  \infer[@{text "(\<or>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<or> Q"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P, Q"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1245
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1246
  This breaks down some disjunction on the right side, replacing it by
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1247
  both disjuncts.  Thus, the sequent calculus is a kind of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1248
  multiple-conclusion logic.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1249
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1250
  To illustrate the use of multiple formulae on the right, let us
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1251
  prove the classical theorem @{text "(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}.  Working
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1252
  backwards, we reduce this formula to a basic sequent:
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1253
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1254
  \infer[@{text "(\<or>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1255
    {\infer[@{text "(\<longrightarrow>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q), (Q \<longrightarrow> P)"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1256
      {\infer[@{text "(\<longrightarrow>R)"}]{@{text "P \<turnstile> Q, (Q \<longrightarrow> P)"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1257
        {@{text "P, Q \<turnstile> Q, P"}}}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1258
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1259
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1260
  This example is typical of the sequent calculus: start with the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1261
  desired theorem and apply rules backwards in a fairly arbitrary
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1262
  manner.  This yields a surprisingly effective proof procedure.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1263
  Quantifiers add only few complications, since Isabelle handles
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 58310
diff changeset
  1264
  parameters and schematic variables.  See @{cite \<open>Chapter 10\<close>
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1265
  "paulson-ml2"} for further discussion.\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1266
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1267
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1268
subsubsection \<open>Simulating sequents by natural deduction\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1269
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1270
text \<open>Isabelle can represent sequents directly, as in the
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1271
  object-logic LK.  But natural deduction is easier to work with, and
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1272
  most object-logics employ it.  Fortunately, we can simulate the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1273
  sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} by the Isabelle formula
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1274
  @{text "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>2 \<Longrightarrow> ... \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> Q\<^sub>1"} where the order of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1275
  the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1276
  Elim-resolution plays a key role in simulating sequent proofs.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1277
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1278
  We can easily handle reasoning on the left.  Elim-resolution with
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1279
  the rules @{text "(\<or>E)"}, @{text "(\<bottom>E)"} and @{text "(\<exists>E)"} achieves
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1280
  a similar effect as the corresponding sequent rules.  For the other
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1281
  connectives, we use sequent-style elimination rules instead of
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1282
  destruction rules such as @{text "(\<and>E1, 2)"} and @{text "(\<forall>E)"}.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1283
  But note that the rule @{text "(\<not>L)"} has no effect under our
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1284
  representation of sequents!
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1285
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1286
  \infer[@{text "(\<not>L)"}]{@{text "\<not> P, \<Gamma> \<turnstile> \<Delta>"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1287
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1288
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1289
  What about reasoning on the right?  Introduction rules can only
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1290
  affect the formula in the conclusion, namely @{text "Q\<^sub>1"}.  The
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1291
  other right-side formulae are represented as negated assumptions,
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1292
  @{text "\<not> Q\<^sub>2, \<dots>, \<not> Q\<^sub>n"}.  In order to operate on one of these, it
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1293
  must first be exchanged with @{text "Q\<^sub>1"}.  Elim-resolution with the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1294
  @{text swap} rule has this effect: @{text "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1295
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1296
  To ensure that swaps occur only when necessary, each introduction
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1297
  rule is converted into a swapped form: it is resolved with the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1298
  second premise of @{text "(swap)"}.  The swapped form of @{text
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1299
  "(\<and>I)"}, which might be called @{text "(\<not>\<and>E)"}, is
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1300
  @{text [display] "\<not> (P \<and> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (\<not> R \<Longrightarrow> Q) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1301
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1302
  Similarly, the swapped form of @{text "(\<longrightarrow>I)"} is
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1303
  @{text [display] "\<not> (P \<longrightarrow> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P \<Longrightarrow> Q) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1304
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1305
  Swapped introduction rules are applied using elim-resolution, which
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1306
  deletes the negated formula.  Our representation of sequents also
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1307
  requires the use of ordinary introduction rules.  If we had no
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1308
  regard for readability of intermediate goal states, we could treat
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1309
  the right side more uniformly by representing sequents as @{text
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1310
  [display] "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> \<bottom>"}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1311
\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1312
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1313
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1314
subsubsection \<open>Extra rules for the sequent calculus\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1315
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1316
text \<open>As mentioned, destruction rules such as @{text "(\<and>E1, 2)"} and
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1317
  @{text "(\<forall>E)"} must be replaced by sequent-style elimination rules.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1318
  In addition, we need rules to embody the classical equivalence
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1319
  between @{text "P \<longrightarrow> Q"} and @{text "\<not> P \<or> Q"}.  The introduction
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1320
  rules @{text "(\<or>I1, 2)"} are replaced by a rule that simulates
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1321
  @{text "(\<or>R)"}: @{text [display] "(\<not> Q \<Longrightarrow> P) \<Longrightarrow> P \<or> Q"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1322
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1323
  The destruction rule @{text "(\<longrightarrow>E)"} is replaced by @{text [display]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1324
  "(P \<longrightarrow> Q) \<Longrightarrow> (\<not> P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1325
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1326
  Quantifier replication also requires special rules.  In classical
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1327
  logic, @{text "\<exists>x. P x"} is equivalent to @{text "\<not> (\<forall>x. \<not> P x)"};
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1328
  the rules @{text "(\<exists>R)"} and @{text "(\<forall>L)"} are dual:
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1329
  \[
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1330
  \infer[@{text "(\<exists>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x"}}{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x, P t"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1331
  \qquad
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1332
  \infer[@{text "(\<forall>L)"}]{@{text "\<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}{@{text "P t, \<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1333
  \]
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1334
  Thus both kinds of quantifier may be replicated.  Theorems requiring
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1335
  multiple uses of a universal formula are easy to invent; consider
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1336
  @{text [display] "(\<forall>x. P x \<longrightarrow> P (f x)) \<and> P a \<longrightarrow> P (f\<^sup>n a)"} for any
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1337
  @{text "n > 1"}.  Natural examples of the multiple use of an
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1338
  existential formula are rare; a standard one is @{text "\<exists>x. \<forall>y. P x
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1339
  \<longrightarrow> P y"}.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1340
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1341
  Forgoing quantifier replication loses completeness, but gains
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1342
  decidability, since the search space becomes finite.  Many useful
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1343
  theorems can be proved without replication, and the search generally
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1344
  delivers its verdict in a reasonable time.  To adopt this approach,
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1345
  represent the sequent rules @{text "(\<exists>R)"}, @{text "(\<exists>L)"} and
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1346
  @{text "(\<forall>R)"} by @{text "(\<exists>I)"}, @{text "(\<exists>E)"} and @{text "(\<forall>I)"},
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1347
  respectively, and put @{text "(\<forall>E)"} into elimination form: @{text
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1348
  [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1349
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1350
  Elim-resolution with this rule will delete the universal formula
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1351
  after a single use.  To replicate universal quantifiers, replace the
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1352
  rule by @{text [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1353
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1354
  To replicate existential quantifiers, replace @{text "(\<exists>I)"} by
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1355
  @{text [display] "(\<not> (\<exists>x. P x) \<Longrightarrow> P t) \<Longrightarrow> \<exists>x. P x"}
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1356
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1357
  All introduction rules mentioned above are also useful in swapped
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1358
  form.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1359
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1360
  Replication makes the search space infinite; we must apply the rules
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1361
  with care.  The classical reasoner distinguishes between safe and
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1362
  unsafe rules, applying the latter only when there is no alternative.
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1363
  Depth-first search may well go down a blind alley; best-first search
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1364
  is better behaved in an infinite search space.  However, quantifier
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1365
  replication is too expensive to prove any but the simplest theorems.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1366
\<close>
42927
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1367
c40adab7568e moved/updated introduction to Classical Reasoner;
wenzelm
parents: 42925
diff changeset
  1368
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1369
subsection \<open>Rule declarations\<close>
42928
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1370
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1371
text \<open>The proof tools of the Classical Reasoner depend on
42928
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1372
  collections of rules declared in the context, which are classified
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1373
  as introduction, elimination or destruction and as \emph{safe} or
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1374
  \emph{unsafe}.  In general, safe rules can be attempted blindly,
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1375
  while unsafe rules must be used with care.  A safe rule must never
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1376
  reduce a provable goal to an unprovable set of subgoals.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1377
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1378
  The rule @{text "P \<Longrightarrow> P \<or> Q"} is unsafe because it reduces @{text "P
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1379
  \<or> Q"} to @{text "P"}, which might turn out as premature choice of an
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1380
  unprovable subgoal.  Any rule is unsafe whose premises contain new
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1381
  unknowns.  The elimination rule @{text "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"} is
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1382
  unsafe, since it is applied via elim-resolution, which discards the
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1383
  assumption @{text "\<forall>x. P x"} and replaces it by the weaker
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1384
  assumption @{text "P t"}.  The rule @{text "P t \<Longrightarrow> \<exists>x. P x"} is
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1385
  unsafe for similar reasons.  The quantifier duplication rule @{text
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1386
  "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"} is unsafe in a different sense:
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1387
  since it keeps the assumption @{text "\<forall>x. P x"}, it is prone to
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1388
  looping.  In classical first-order logic, all rules are safe except
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1389
  those mentioned above.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1390
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1391
  The safe~/ unsafe distinction is vague, and may be regarded merely
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1392
  as a way of giving some rules priority over others.  One could argue
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1393
  that @{text "(\<or>E)"} is unsafe, because repeated application of it
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1394
  could generate exponentially many subgoals.  Induction rules are
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1395
  unsafe because inductive proofs are difficult to set up
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1396
  automatically.  Any inference is unsafe that instantiates an unknown
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1397
  in the proof state --- thus matching must be used, rather than
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1398
  unification.  Even proof by assumption is unsafe if it instantiates
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1399
  unknowns shared with other subgoals.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1400
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1401
  \begin{matharray}{rcl}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1402
    @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1403
    @{attribute_def intro} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1404
    @{attribute_def elim} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1405
    @{attribute_def dest} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1406
    @{attribute_def rule} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1407
    @{attribute_def iff} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1408
    @{attribute_def swapped} & : & @{text attribute} \\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1409
  \end{matharray}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1410
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1411
  @{rail \<open>
42928
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1412
    (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}?
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1413
    ;
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1414
    @@{attribute rule} 'del'
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1415
    ;
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1416
    @@{attribute iff} (((() | 'add') '?'?) | 'del')
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1417
  \<close>}
42928
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1418
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1419
  \begin{description}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1420
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1421
  \item @{command "print_claset"} prints the collection of rules
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1422
  declared to the Classical Reasoner, i.e.\ the @{ML_type claset}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1423
  within the context.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1424
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1425
  \item @{attribute intro}, @{attribute elim}, and @{attribute dest}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1426
  declare introduction, elimination, and destruction rules,
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1427
  respectively.  By default, rules are considered as \emph{unsafe}
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1428
  (i.e.\ not applied blindly without backtracking), while ``@{text
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1429
  "!"}'' classifies as \emph{safe}.  Rule declarations marked by
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1430
  ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1431
  \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1432
  of the @{method rule} method).  The optional natural number
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1433
  specifies an explicit weight argument, which is ignored by the
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1434
  automated reasoning tools, but determines the search order of single
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1435
  rule steps.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1436
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1437
  Introduction rules are those that can be applied using ordinary
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1438
  resolution.  Their swapped forms are generated internally, which
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1439
  will be applied using elim-resolution.  Elimination rules are
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1440
  applied using elim-resolution.  Rules are sorted by the number of
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1441
  new subgoals they will yield; rules that generate the fewest
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1442
  subgoals will be tried first.  Otherwise, later declarations take
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1443
  precedence over earlier ones.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1444
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1445
  Rules already present in the context with the same classification
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1446
  are ignored.  A warning is printed if the rule has already been
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1447
  added with some other classification, but the rule is added anyway
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1448
  as requested.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1449
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1450
  \item @{attribute rule}~@{text del} deletes all occurrences of a
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1451
  rule from the classical context, regardless of its classification as
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1452
  introduction~/ elimination~/ destruction and safe~/ unsafe.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1453
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1454
  \item @{attribute iff} declares logical equivalences to the
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1455
  Simplifier and the Classical reasoner at the same time.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1456
  Non-conditional rules result in a safe introduction and elimination
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1457
  pair; conditional ones are considered unsafe.  Rules with negative
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1458
  conclusion are automatically inverted (using @{text "\<not>"}-elimination
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1459
  internally).
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1460
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1461
  The ``@{text "?"}'' version of @{attribute iff} declares rules to
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1462
  the Isabelle/Pure context only, and omits the Simplifier
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1463
  declaration.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1464
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1465
  \item @{attribute swapped} turns an introduction rule into an
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1466
  elimination, by resolving with the classical swap principle @{text
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1467
  "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"} in the second position.  This is mainly for
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1468
  illustrative purposes: the Classical Reasoner already swaps rules
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1469
  internally as explained above.
9d946de41120 updated and re-unified classical rule declarations;
wenzelm
parents: 42927
diff changeset
  1470
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1471
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1472
\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1473
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1474
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1475
subsection \<open>Structured methods\<close>
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1476
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1477
text \<open>
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1478
  \begin{matharray}{rcl}
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1479
    @{method_def rule} & : & @{text method} \\
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1480
    @{method_def contradiction} & : & @{text method} \\
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1481
  \end{matharray}
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1482
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1483
  @{rail \<open>
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1484
    @@{method rule} @{syntax thmrefs}?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1485
  \<close>}
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1486
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1487
  \begin{description}
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1488
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1489
  \item @{method rule} as offered by the Classical Reasoner is a
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1490
  refinement over the Pure one (see \secref{sec:pure-meth-att}).  Both
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1491
  versions work the same, but the classical version observes the
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1492
  classical rule context in addition to that of Isabelle/Pure.
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1493
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1494
  Common object logics (HOL, ZF, etc.) declare a rich collection of
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1495
  classical rules (even if these would qualify as intuitionistic
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1496
  ones), but only few declarations to the rule context of
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1497
  Isabelle/Pure (\secref{sec:pure-meth-att}).
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1498
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1499
  \item @{method contradiction} solves some goal by contradiction,
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1500
  deriving any result from both @{text "\<not> A"} and @{text A}.  Chained
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1501
  facts, which are guaranteed to participate, may appear in either
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1502
  order.
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1503
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1504
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1505
\<close>
43365
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1506
f8944cb2468f tuned sections;
wenzelm
parents: 43332
diff changeset
  1507
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1508
subsection \<open>Fully automated methods\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1509
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1510
text \<open>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1511
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1512
    @{method_def blast} & : & @{text method} \\
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1513
    @{method_def auto} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1514
    @{method_def force} & : & @{text method} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1515
    @{method_def fast} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1516
    @{method_def slow} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1517
    @{method_def best} & : & @{text method} \\
44911
884d27ede438 fastsimp -> fastforce in doc
nipkow
parents: 44094
diff changeset
  1518
    @{method_def fastforce} & : & @{text method} \\
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1519
    @{method_def slowsimp} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1520
    @{method_def bestsimp} & : & @{text method} \\
43367
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1521
    @{method_def deepen} & : & @{text method} \\
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1522
  \end{matharray}
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1523
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1524
  @{rail \<open>
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1525
    @@{method blast} @{syntax nat}? (@{syntax clamod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1526
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1527
    @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1528
    ;
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1529
    @@{method force} (@{syntax clasimpmod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1530
    ;
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1531
    (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * )
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1532
    ;
44911
884d27ede438 fastsimp -> fastforce in doc
nipkow
parents: 44094
diff changeset
  1533
    (@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp})
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1534
      (@{syntax clasimpmod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1535
    ;
43367
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1536
    @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * )
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1537
    ;
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1538
    @{syntax_def clamod}:
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1539
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1540
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1541
    @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') |
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1542
      ('cong' | 'split') (() | 'add' | 'del') |
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1543
      'iff' (((() | 'add') '?'?) | 'del') |
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1544
      (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs}
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1545
  \<close>}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1546
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1547
  \begin{description}
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1548
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1549
  \item @{method blast} is a separate classical tableau prover that
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1550
  uses the same classical rule declarations as explained before.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1551
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1552
  Proof search is coded directly in ML using special data structures.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1553
  A successful proof is then reconstructed using regular Isabelle
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1554
  inferences.  It is faster and more powerful than the other classical
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1555
  reasoning tools, but has major limitations too.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1556
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1557
  \begin{itemize}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1558
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1559
  \item It does not use the classical wrapper tacticals, such as the
44911
884d27ede438 fastsimp -> fastforce in doc
nipkow
parents: 44094
diff changeset
  1560
  integration with the Simplifier of @{method fastforce}.
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1561
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1562
  \item It does not perform higher-order unification, as needed by the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1563
  rule @{thm [source=false] rangeI} in HOL.  There are often
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1564
  alternatives to such rules, for example @{thm [source=false]
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1565
  range_eqI}.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1566
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1567
  \item Function variables may only be applied to parameters of the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1568
  subgoal.  (This restriction arises because the prover does not use
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1569
  higher-order unification.)  If other function variables are present
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1570
  then the prover will fail with the message \texttt{Function Var's
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1571
  argument not a bound variable}.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1572
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1573
  \item Its proof strategy is more general than @{method fast} but can
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1574
  be slower.  If @{method blast} fails or seems to be running forever,
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1575
  try @{method fast} and the other proof tools described below.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1576
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1577
  \end{itemize}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1578
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1579
  The optional integer argument specifies a bound for the number of
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1580
  unsafe steps used in a proof.  By default, @{method blast} starts
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1581
  with a bound of 0 and increases it successively to 20.  In contrast,
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1582
  @{text "(blast lim)"} tries to prove the goal using a search bound
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1583
  of @{text "lim"}.  Sometimes a slow proof using @{method blast} can
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1584
  be made much faster by supplying the successful search bound to this
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1585
  proof method instead.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1586
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1587
  \item @{method auto} combines classical reasoning with
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1588
  simplification.  It is intended for situations where there are a lot
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1589
  of mostly trivial subgoals; it proves all the easy ones, leaving the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1590
  ones it cannot prove.  Occasionally, attempting to prove the hard
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1591
  ones may take a long time.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1592
43332
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
  1593
  The optional depth arguments in @{text "(auto m n)"} refer to its
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
  1594
  builtin classical reasoning procedures: @{text m} (default 4) is for
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
  1595
  @{method blast}, which is tried first, and @{text n} (default 2) is
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
  1596
  for a slower but more general alternative that also takes wrappers
dca2c7c598f0 document depth arguments of method "auto";
wenzelm
parents: 42930
diff changeset
  1597
  into account.
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1598
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1599
  \item @{method force} is intended to prove the first subgoal
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1600
  completely, using many fancy proof tools and performing a rather
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1601
  exhaustive search.  As a result, proof attempts may take rather long
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1602
  or diverge easily.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1603
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1604
  \item @{method fast}, @{method best}, @{method slow} attempt to
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1605
  prove the first subgoal using sequent-style reasoning as explained
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1606
  before.  Unlike @{method blast}, they construct proofs directly in
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1607
  Isabelle.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1608
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1609
  There is a difference in search strategy and back-tracking: @{method
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1610
  fast} uses depth-first search and @{method best} uses best-first
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1611
  search (guided by a heuristic function: normally the total size of
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1612
  the proof state).
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1613
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1614
  Method @{method slow} is like @{method fast}, but conducts a broader
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1615
  search: it may, when backtracking from a failed proof attempt, undo
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1616
  even the step of proving a subgoal by assumption.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1617
47967
c422128d3889 discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents: 47497
diff changeset
  1618
  \item @{method fastforce}, @{method slowsimp}, @{method bestsimp}
c422128d3889 discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents: 47497
diff changeset
  1619
  are like @{method fast}, @{method slow}, @{method best},
c422128d3889 discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents: 47497
diff changeset
  1620
  respectively, but use the Simplifier as additional wrapper. The name
c422128d3889 discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents: 47497
diff changeset
  1621
  @{method fastforce}, reflects the behaviour of this popular method
c422128d3889 discontinued obsolete method fastsimp / tactic fast_simp_tac;
wenzelm
parents: 47497
diff changeset
  1622
  better without requiring an understanding of its implementation.
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1623
43367
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1624
  \item @{method deepen} works by exhaustive search up to a certain
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1625
  depth.  The start depth is 4 (unless specified explicitly), and the
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1626
  depth is increased iteratively up to 10.  Unsafe rules are modified
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1627
  to preserve the formula they act on, so that it be used repeatedly.
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1628
  This method can prove more goals than @{method fast}, but is much
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1629
  slower, for example if the assumptions have many universal
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1630
  quantifiers.
3f1469de9e47 cover method "deepen" concisely;
wenzelm
parents: 43366
diff changeset
  1631
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1632
  \end{description}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1633
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1634
  Any of the above methods support additional modifiers of the context
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1635
  of classical (and simplifier) rules, but the ones related to the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1636
  Simplifier are explicitly prefixed by @{text simp} here.  The
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1637
  semantics of these ad-hoc rule declarations is analogous to the
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1638
  attributes given before.  Facts provided by forward chaining are
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1639
  inserted into the goal before commencing proof search.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1640
\<close>
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1641
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1642
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1643
subsection \<open>Partially automated methods\<close>
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1644
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1645
text \<open>These proof methods may help in situations when the
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1646
  fully-automated tools fail.  The result is a simpler subgoal that
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1647
  can be tackled by other means, such as by manual instantiation of
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1648
  quantifiers.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1649
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1650
  \begin{matharray}{rcl}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1651
    @{method_def safe} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1652
    @{method_def clarify} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1653
    @{method_def clarsimp} & : & @{text method} \\
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1654
  \end{matharray}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1655
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1656
  @{rail \<open>
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1657
    (@@{method safe} | @@{method clarify}) (@{syntax clamod} * )
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1658
    ;
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1659
    @@{method clarsimp} (@{syntax clasimpmod} * )
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1660
  \<close>}
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1661
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1662
  \begin{description}
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1663
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1664
  \item @{method safe} repeatedly performs safe steps on all subgoals.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1665
  It is deterministic, with at most one outcome.
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1666
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1667
  \item @{method clarify} performs a series of safe steps without
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1668
  splitting subgoals; see also @{method clarify_step}.
42930
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1669
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1670
  \item @{method clarsimp} acts like @{method clarify}, but also does
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1671
  simplification.  Note that if the Simplifier context includes a
41394a61cca9 updated and re-unified classical proof methods;
wenzelm
parents: 42929
diff changeset
  1672
  splitter for the premises, the subgoal may still be split.
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1673
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1674
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1675
\<close>
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1676
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1677
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1678
subsection \<open>Single-step tactics\<close>
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1679
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1680
text \<open>
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1681
  \begin{matharray}{rcl}
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1682
    @{method_def safe_step} & : & @{text method} \\
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1683
    @{method_def inst_step} & : & @{text method} \\
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1684
    @{method_def step} & : & @{text method} \\
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1685
    @{method_def slow_step} & : & @{text method} \\
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1686
    @{method_def clarify_step} & : &  @{text method} \\
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1687
  \end{matharray}
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1688
50070
e447ad4d6edd avoid clash of terminology wrt. "semi-automated" in the sense of Isar (e.g. method "rule");
wenzelm
parents: 50068
diff changeset
  1689
  These are the primitive tactics behind the automated proof methods
e447ad4d6edd avoid clash of terminology wrt. "semi-automated" in the sense of Isar (e.g. method "rule");
wenzelm
parents: 50068
diff changeset
  1690
  of the Classical Reasoner.  By calling them yourself, you can
e447ad4d6edd avoid clash of terminology wrt. "semi-automated" in the sense of Isar (e.g. method "rule");
wenzelm
parents: 50068
diff changeset
  1691
  execute these procedures one step at a time.
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1692
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1693
  \begin{description}
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1694
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1695
  \item @{method safe_step} performs a safe step on the first subgoal.
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1696
  The safe wrapper tacticals are applied to a tactic that may include
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1697
  proof by assumption or Modus Ponens (taking care not to instantiate
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1698
  unknowns), or substitution.
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1699
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1700
  \item @{method inst_step} is like @{method safe_step}, but allows
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1701
  unknowns to be instantiated.
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1702
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1703
  \item @{method step} is the basic step of the proof procedure, it
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1704
  operates on the first subgoal.  The unsafe wrapper tacticals are
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1705
  applied to a tactic that tries @{method safe}, @{method inst_step},
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1706
  or applies an unsafe rule from the context.
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1707
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1708
  \item @{method slow_step} resembles @{method step}, but allows
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1709
  backtracking between using safe rules with instantiation (@{method
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1710
  inst_step}) and using unsafe rules.  The resulting search space is
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1711
  larger.
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1712
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1713
  \item @{method clarify_step} performs a safe step on the first
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1714
  subgoal; no splitting step is applied.  For example, the subgoal
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1715
  @{text "A \<and> B"} is left as a conjunction.  Proof by assumption,
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1716
  Modus Ponens, etc., may be performed provided they do not
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1717
  instantiate unknowns.  Assumptions of the form @{text "x = t"} may
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1718
  be eliminated.  The safe wrapper tactical is applied.
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1719
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1720
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1721
\<close>
43366
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1722
9cbcab5c780a moved/updated single-step tactics;
wenzelm
parents: 43365
diff changeset
  1723
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1724
subsection \<open>Modifying the search step\<close>
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1725
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1726
text \<open>
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1727
  \begin{mldecls}
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1728
    @{index_ML_type wrapper: "(int -> tactic) -> (int -> tactic)"} \\[0.5ex]
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1729
    @{index_ML_op addSWrapper: "Proof.context *
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1730
  (string * (Proof.context -> wrapper)) -> Proof.context"} \\
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1731
    @{index_ML_op addSbefore: "Proof.context *
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1732
  (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1733
    @{index_ML_op addSafter: "Proof.context *
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1734
  (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1735
    @{index_ML_op delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1736
    @{index_ML_op addWrapper: "Proof.context *
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1737
  (string * (Proof.context -> wrapper)) -> Proof.context"} \\
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1738
    @{index_ML_op addbefore: "Proof.context *
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1739
  (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1740
    @{index_ML_op addafter: "Proof.context *
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1741
  (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1742
    @{index_ML_op delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1743
    @{index_ML addSss: "Proof.context -> Proof.context"} \\
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1744
    @{index_ML addss: "Proof.context -> Proof.context"} \\
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1745
  \end{mldecls}
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1746
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1747
  The proof strategy of the Classical Reasoner is simple.  Perform as
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1748
  many safe inferences as possible; or else, apply certain safe rules,
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1749
  allowing instantiation of unknowns; or else, apply an unsafe rule.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1750
  The tactics also eliminate assumptions of the form @{text "x = t"}
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1751
  by substitution if they have been set up to do so.  They may perform
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1752
  a form of Modus Ponens: if there are assumptions @{text "P \<longrightarrow> Q"} and
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1753
  @{text "P"}, then replace @{text "P \<longrightarrow> Q"} by @{text "Q"}.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1754
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1755
  The classical reasoning tools --- except @{method blast} --- allow
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1756
  to modify this basic proof strategy by applying two lists of
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1757
  arbitrary \emph{wrapper tacticals} to it.  The first wrapper list,
50108
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1758
  which is considered to contain safe wrappers only, affects @{method
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1759
  safe_step} and all the tactics that call it.  The second one, which
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1760
  may contain unsafe wrappers, affects the unsafe parts of @{method
f171b5240c31 method setup for Classical steps;
wenzelm
parents: 50083
diff changeset
  1761
  step}, @{method slow_step}, and the tactics that call them.  A
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1762
  wrapper transforms each step of the search, for example by
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1763
  attempting other tactics before or after the original step tactic.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1764
  All members of a wrapper list are applied in turn to the respective
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1765
  step tactic.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1766
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1767
  Initially the two wrapper lists are empty, which means no
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1768
  modification of the step tactics. Safe and unsafe wrappers are added
59905
678c9e625782 misc tuning -- keep name space more clean;
wenzelm
parents: 59853
diff changeset
  1769
  to the context with the functions given below, supplying them with
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1770
  wrapper names.  These names may be used to selectively delete
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1771
  wrappers.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1772
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1773
  \begin{description}
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1774
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1775
  \item @{text "ctxt addSWrapper (name, wrapper)"} adds a new wrapper,
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1776
  which should yield a safe tactic, to modify the existing safe step
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1777
  tactic.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1778
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1779
  \item @{text "ctxt addSbefore (name, tac)"} adds the given tactic as a
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1780
  safe wrapper, such that it is tried \emph{before} each safe step of
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1781
  the search.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1782
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1783
  \item @{text "ctxt addSafter (name, tac)"} adds the given tactic as a
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1784
  safe wrapper, such that it is tried when a safe step of the search
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1785
  would fail.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1786
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1787
  \item @{text "ctxt delSWrapper name"} deletes the safe wrapper with
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1788
  the given name.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1789
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1790
  \item @{text "ctxt addWrapper (name, wrapper)"} adds a new wrapper to
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1791
  modify the existing (unsafe) step tactic.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1792
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1793
  \item @{text "ctxt addbefore (name, tac)"} adds the given tactic as an
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1794
  unsafe wrapper, such that it its result is concatenated
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1795
  \emph{before} the result of each unsafe step.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1796
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1797
  \item @{text "ctxt addafter (name, tac)"} adds the given tactic as an
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1798
  unsafe wrapper, such that it its result is concatenated \emph{after}
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1799
  the result of each unsafe step.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1800
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 50108
diff changeset
  1801
  \item @{text "ctxt delWrapper name"} deletes the unsafe wrapper with
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1802
  the given name.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1803
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1804
  \item @{text "addSss"} adds the simpset of the context to its
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1805
  classical set. The assumptions and goal will be simplified, in a
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1806
  rather safe way, after each safe step of the search.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1807
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1808
  \item @{text "addss"} adds the simpset of the context to its
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1809
  classical set. The assumptions and goal will be simplified, before
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1810
  the each unsafe step of the search.
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1811
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1812
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1813
\<close>
50071
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1814
959548c3b947 moved classical wrappers to IsarRef;
wenzelm
parents: 50070
diff changeset
  1815
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1816
section \<open>Object-logic setup \label{sec:object-logic}\<close>
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1817
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1818
text \<open>
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1819
  \begin{matharray}{rcl}
28761
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1820
    @{command_def "judgment"} & : & @{text "theory \<rightarrow> theory"} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1821
    @{method_def atomize} & : & @{text method} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1822
    @{attribute_def atomize} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1823
    @{attribute_def rule_format} & : & @{text attribute} \\
9ec4482c9201 updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents: 28760
diff changeset
  1824
    @{attribute_def rulify} & : & @{text attribute} \\
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1825
  \end{matharray}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1826
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1827
  The very starting point for any Isabelle object-logic is a ``truth
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1828
  judgment'' that links object-level statements to the meta-logic
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1829
  (with its minimal language of @{text prop} that covers universal
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1830
  quantification @{text "\<And>"} and implication @{text "\<Longrightarrow>"}).
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1831
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1832
  Common object-logics are sufficiently expressive to internalize rule
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1833
  statements over @{text "\<And>"} and @{text "\<Longrightarrow>"} within their own
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1834
  language.  This is useful in certain situations where a rule needs
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1835
  to be viewed as an atomic statement from the meta-level perspective,
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1836
  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
  1837
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1838
  From the following language elements, only the @{method atomize}
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1839
  method and @{attribute rule_format} attribute are occasionally
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1840
  required by end-users, the rest is for those who need to setup their
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1841
  own object-logic.  In the latter case existing formulations of
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1842
  Isabelle/FOL or Isabelle/HOL may be taken as realistic examples.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1843
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1844
  Generic tools may refer to the information provided by object-logic
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1845
  declarations internally.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1846
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1847
  @{rail \<open>
46494
ea2ae63336f3 clarified outer syntax "constdecl", which is only local to some rail diagrams;
wenzelm
parents: 46277
diff changeset
  1848
    @@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}?
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1849
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1850
    @@{attribute atomize} ('(' 'full' ')')?
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1851
    ;
42596
6c621a9d612a modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents: 40291
diff changeset
  1852
    @@{attribute rule_format} ('(' 'noasm' ')')?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1853
  \<close>}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1854
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1855
  \begin{description}
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1856
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1857
  \item @{command "judgment"}~@{text "c :: \<sigma> (mx)"} declares constant
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1858
  @{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
  1859
  type @{text \<sigma>} should specify a coercion of the category of
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1860
  object-level propositions to @{text prop} of the Pure meta-logic;
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1861
  the mixfix annotation @{text "(mx)"} would typically just link the
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1862
  object language (internally of syntactic category @{text logic})
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1863
  with that of @{text prop}.  Only one @{command "judgment"}
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1864
  declaration may be given in any theory development.
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1865
  
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1866
  \item @{method atomize} (as a method) rewrites any non-atomic
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1867
  premises of a sub-goal, using the meta-level equations declared via
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1868
  @{attribute atomize} (as an attribute) beforehand.  As a result,
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1869
  heavily nested goals become amenable to fundamental operations such
42626
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 42617
diff changeset
  1870
  as resolution (cf.\ the @{method (Pure) rule} method).  Giving the ``@{text
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1871
  "(full)"}'' option here means to turn the whole subgoal into an
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1872
  object-statement (if possible), including the outermost parameters
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1873
  and assumptions as well.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1874
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1875
  A typical collection of @{attribute atomize} rules for a particular
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1876
  object-logic would provide an internalization for each of the
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1877
  connectives of @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"}.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1878
  Meta-level conjunction should be covered as well (this is
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1879
  particularly important for locales, see \secref{sec:locale}).
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1880
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1881
  \item @{attribute rule_format} rewrites a theorem by the equalities
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1882
  declared as @{attribute rulify} rules in the current object-logic.
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1883
  By default, the result is fully normalized, including assumptions
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1884
  and conclusions at any depth.  The @{text "(no_asm)"} option
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1885
  restricts the transformation to the conclusion of a rule.
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1886
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1887
  In common object-logics (HOL, FOL, ZF), the effect of @{attribute
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1888
  rule_format} is to replace (bounded) universal quantification
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1889
  (@{text "\<forall>"}) and implication (@{text "\<longrightarrow>"}) by the corresponding
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1890
  rule statements over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1891
28760
cbc435f7b16b unified use of declaration environment with IsarImplementation;
wenzelm
parents: 28754
diff changeset
  1892
  \end{description}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1893
\<close>
26790
e8cc166ba123 converted "General logic setup";
wenzelm
parents: 26789
diff changeset
  1894
50083
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1895
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1896
section \<open>Tracing higher-order unification\<close>
50083
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1897
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1898
text \<open>
50083
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1899
  \begin{tabular}{rcll}
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1900
    @{attribute_def unify_trace_simp} & : & @{text "attribute"} & default @{text "false"} \\
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1901
    @{attribute_def unify_trace_types} & : & @{text "attribute"} & default @{text "false"} \\
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1902
    @{attribute_def unify_trace_bound} & : & @{text "attribute"} & default @{text "50"} \\
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1903
    @{attribute_def unify_search_bound} & : & @{text "attribute"} & default @{text "60"} \\
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1904
  \end{tabular}
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1905
  \medskip
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1906
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1907
  Higher-order unification works well in most practical situations,
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1908
  but sometimes needs extra care to identify problems.  These tracing
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1909
  options may help.
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1910
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1911
  \begin{description}
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1912
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1913
  \item @{attribute unify_trace_simp} controls tracing of the
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1914
  simplification phase of higher-order unification.
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1915
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1916
  \item @{attribute unify_trace_types} controls warnings of
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1917
  incompleteness, when unification is not considering all possible
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1918
  instantiations of schematic type variables.
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1919
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1920
  \item @{attribute unify_trace_bound} determines the depth where
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1921
  unification starts to print tracing information once it reaches
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1922
  depth; 0 for full tracing.  At the default value, tracing
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1923
  information is almost never printed in practice.
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1924
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1925
  \item @{attribute unify_search_bound} prevents unification from
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1926
  searching past the given depth.  Because of this bound, higher-order
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1927
  unification cannot return an infinite sequence, though it can return
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1928
  an exponentially long one.  The search rarely approaches the default
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1929
  value in practice.  If the search is cut off, unification prints a
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1930
  warning ``Unification bound exceeded''.
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1931
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1932
  \end{description}
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1933
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1934
  \begin{warn}
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1935
  Options for unification cannot be modified in a local context.  Only
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1936
  the global theory content is taken into account.
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1937
  \end{warn}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1938
\<close>
50083
01605e79c569 updated unification options;
wenzelm
parents: 50080
diff changeset
  1939
26782
19363c70b5c4 converted generic.tex to Thy/Generic.thy;
wenzelm
parents:
diff changeset
  1940
end