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