CookBook/Tactical.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 15 Feb 2009 18:58:21 +0000
changeset 120 c39f83d8daeb
parent 118 5f003fdf2653
child 126 fcc0e6e54dca
permissions -rw-r--r--
some polishing; split up the file External Solver into two
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Tactical
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
     2
imports Base FirstSteps
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
chapter {* Tactical Reasoning\label{chp:tactical} *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
     8
  The main reason for descending to the ML-level of Isabelle is to be able to
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
     9
  implement automatic proof procedures. Such proof procedures usually lessen
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    10
  considerably the burden of manual reasoning, for example, when introducing
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    11
  new definitions. These proof procedures are centred around refining a goal
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    12
  state using tactics. This is similar to the \isacommand{apply}-style
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    13
  reasoning at the user level, where goals are modified in a sequence of proof
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    14
  steps until all of them are solved. However, there are also more structured
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    15
  operations available on the ML-level that help with the handling of
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    16
  variables and assumptions.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    17
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    20
section {* Basics of Reasoning with Tactics*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
text {*
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    23
  To see how tactics work, let us first transcribe a simple \isacommand{apply}-style proof 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
    24
  into ML. Suppose the following proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
apply(erule disjE)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
apply(rule disjI2)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
apply(rule disjI1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  This proof translates to the following ML-code.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
@{ML_response_fake [display,gray]
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
"let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  Goal.prove ctxt [\"P\", \"Q\"] [] goal 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
   (fn _ => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
      etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
      THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
      THEN atac 1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  To start the proof, the function @{ML "Goal.prove"}~@{text "ctxt xs As C
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    53
  tac"} sets up a goal state for proving the goal @{text C} 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    54
  (that is @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    55
  assumptions @{text As} (happens to be empty) with the variables
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  @{text xs} that will be generalised once the goal is proved (in our case
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  @{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  it can make use of the local assumptions (there are none in this example). 
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    59
  The functions @{ML etac}, @{ML rtac} and @{ML atac} in the code above correspond to 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  @{text erule}, @{text rule} and @{text assumption}, respectively. 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    61
  The operator @{ML THEN} strings the tactics together. 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  \begin{readmore}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    64
  To learn more about the function @{ML Goal.prove} see \isccite{sec:results}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
    65
  and the file @{ML_file "Pure/goal.ML"}.  See @{ML_file
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    66
  "Pure/tactic.ML"} and @{ML_file "Pure/tctical.ML"} for the code of basic
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    67
  tactics and tactic combinators; see also Chapters 3 and 4 in the old
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    68
  Isabelle Reference Manual.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    71
  Note that in the code above we use antiquotations for referencing the theorems. Many theorems
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    72
  also have ML-bindings with the same name. Therefore, we could also just have
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
    73
  written @{ML "etac disjE 1"}, or in case where there are no ML-binding obtain
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    74
  the theorem dynamically using the function @{ML thm}; for example 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
    75
  \mbox{@{ML  "etac (thm \"disjE\") 1"}}. Both ways however are considered bad style! 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    76
  The reason
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    77
  is that the binding for @{ML disjE} can be re-assigned by the user and thus
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    78
  one does not have complete control over which theorem is actually
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    79
  applied. This problem is nicely prevented by using antiquotations, because
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    80
  then the theorems are fixed statically at compile-time.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    82
  During the development of automatic proof procedures, you will often find it 
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    83
  necessary to test a tactic on examples. This can be conveniently 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  Consider the following sequence of tactics
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
ML{*val foo_tac = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
      (etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
       THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
       THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
       THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
       THEN atac 1)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
text {* and the Isabelle proof: *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
lemma "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
apply(tactic {* foo_tac *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   102
  By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   103
  user level of Isabelle the tactic @{ML foo_tac} or 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   104
  any other function that returns a tactic.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   106
  The tactic @{ML foo_tac} is just a sequence of simple tactics stringed
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   107
  together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   108
  has a hard-coded number that stands for the subgoal analysed by the
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   109
  tactic (@{text "1"} stands for the first, or top-most, subgoal). This hard-coding
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   110
  of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   111
  you can write\label{tac:footacprime}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
ML{*val foo_tac' = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
      (etac @{thm disjE} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
       THEN' rtac @{thm disjI2} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
       THEN' atac 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
       THEN' rtac @{thm disjI1} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
       THEN' atac)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
text {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  and then give the number for the subgoal explicitly when the tactic is
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   123
  called. So in the next proof you can first discharge the second subgoal, and
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   124
  subsequently the first.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
   and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
apply(tactic {* foo_tac' 2 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
text {*
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   134
  This kind of addressing is more difficult to achieve when the goal is 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   135
  hard-coded inside the tactic. For most operators that combine tactics 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   136
  (@{ML THEN} is only one such operator) a ``primed'' version exists.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   137
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   138
  The tactics @{ML foo_tac} and @{ML foo_tac'} are very specific for
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   139
  analysing goals being only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   140
  of this form, then they return the error message:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   141
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   142
  \begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   143
  @{text "*** empty result sequence -- proof command failed"}\\
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   144
  @{text "*** At command \"apply\"."}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   145
  \end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   146
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   147
  This means the tactics failed. The reason for this error message is that tactics 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   148
  are functions mapping a goal state to a (lazy) sequence of successor states. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   149
  Hence the type of a tactic is:
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   150
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   152
ML{*type tactic = thm -> thm Seq.seq*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   154
text {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   155
  By convention, if a tactic fails, then it should return the empty sequence. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   156
  Therefore, if you write your own tactics, they  should not raise exceptions 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   157
  willy-nilly; only in very grave failure situations should a tactic raise the 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   158
  exception @{text THM}.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   159
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   160
  The simplest tactics are @{ML no_tac} and @{ML all_tac}. The first returns
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   161
  the empty sequence and is defined as
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   162
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   163
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   164
ML{*fun no_tac thm = Seq.empty*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   165
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   166
text {*
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   167
  which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   168
  up in a single member sequence; it is defined as
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   169
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   170
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   171
ML{*fun all_tac thm = Seq.single thm*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   172
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   173
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   174
  which means @{ML all_tac} always succeeds, but also does not make any progress 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   175
  with the proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   177
  The lazy list of possible successor goal states shows through at the user-level 
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   178
  of Isabelle when using the command \isacommand{back}. For instance in the 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   179
  following proof there are two possibilities for how to apply 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   180
  @{ML foo_tac'}: either using the first assumption or the second.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
back
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   188
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   190
  By using \isacommand{back}, we construct the proof that uses the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   191
  second assumption. While in the proof above, it does not really matter which 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   192
  assumption is used, in more interesting cases provability might depend
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   193
  on exploring different possibilities.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   194
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   197
  sequences. In day-to-day Isabelle programming, however, one rarely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   198
  constructs sequences explicitly, but uses the predefined tactics and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   199
  tactic combinators instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   202
  It might be surprising that tactics, which transform
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   203
  one goal state to the next, are functions from theorems to theorem 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   204
  (sequences). The surprise resolves by knowing that every 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   205
  goal state is indeed a theorem. To shed more light on this,
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   206
  let us modify the code of @{ML all_tac} to obtain the following
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   207
  tactic
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   208
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   209
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   210
ML{*fun my_print_tac ctxt thm =
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   211
 let
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   212
   val _ = warning (str_of_thm ctxt thm)
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   213
 in 
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   214
   Seq.single thm
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   215
 end*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   216
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   217
text_raw {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   218
  \begin{figure}[p]
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   219
  \begin{isabelle}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   220
*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   221
lemma shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B" 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   222
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   223
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   224
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   225
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   226
      \end{minipage}\medskip      
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   227
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   228
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   229
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   230
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   231
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   232
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   233
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   234
      \end{minipage}\medskip
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   237
apply(rule conjI)
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   238
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   239
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   240
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   241
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   242
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   243
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   244
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   245
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   246
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   247
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   248
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   249
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   250
      \end{minipage}\medskip
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   251
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   252
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   253
apply(assumption)
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   254
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   255
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   256
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   257
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   258
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   259
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   260
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   261
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   262
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   263
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   264
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   265
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   266
      \end{minipage}\medskip
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   267
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   268
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   269
apply(assumption)
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   270
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   271
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   272
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   273
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   274
      \end{minipage}\medskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   275
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   276
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   277
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   278
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   279
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   280
      @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   281
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   282
      \end{minipage}\medskip   
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   283
   *}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   284
done
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   285
text_raw {*  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   286
  \end{isabelle}
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   287
  \caption{The figure shows a proof where each intermediate goal state is
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   288
  printed by the Isabelle system and by @{ML my_print_tac}. The latter shows
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   289
  the goal state as represented internally (highlighted boxes). This
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   290
  illustrates that every goal state in Isabelle is represented by a theorem:
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   291
  when we start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}} the theorem is
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   292
  @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when we finish the proof the
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   293
  theorem is @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}.\label{fig:goalstates}}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   294
  \end{figure}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   295
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   296
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   297
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   298
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   299
  which prints out the given theorem (using the string-function defined in
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   300
  Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   301
  tactic we are in the position to inspect every goal state in a
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   302
  proof. Consider now the proof in Figure~\ref{fig:goalstates}: as can be seen, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   303
  internally every goal state is an implication of the form
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   304
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   305
  @{text[display] "A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> (C)"}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   306
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   307
  where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   308
  the subgoals. So after setting up the lemma, the goal state is always of the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   309
  form @{text "C \<Longrightarrow> (C)"}; when the proof is finished we are left with @{text
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   310
  "(C)"}. Since the goal @{term C} can potentially be an implication, there is
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   311
  a ``protector'' wrapped around it (in from of an outermost constant @{text
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   312
  "Const (\"prop\", bool \<Rightarrow> bool)"}; however this constant
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   313
  is invisible in the figure). This prevents that premises of @{text C} are
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   314
  mis-interpreted as open subgoals. While tactics can operate on the subgoals
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   315
  (the @{text "A\<^isub>i"} above), they are expected to leave the conclusion
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   316
  @{term C} intact, with the exception of possibly instantiating schematic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   317
  variables. If you use the predefined tactics, which we describe in the next
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   318
  section, this will always be the case.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   319
 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   320
  \begin{readmore}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   321
  For more information about the internals of goals see \isccite{sec:tactical-goals}.
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   322
  \end{readmore}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   323
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   324
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   325
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   326
section {* Simple Tactics *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   328
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   329
  Let us start with the tactic @{ML print_tac}, which is quite useful for low-level 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   330
  debugging of tactics. It just prints out a message and the current goal state. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   331
  Processing the proof
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   332
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   333
 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   334
lemma shows "False \<Longrightarrow> True"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   335
apply(tactic {* print_tac "foo message" *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   336
txt{*gives:\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   337
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   338
     \begin{minipage}{\textwidth}\small
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   339
     @{text "foo message"}\\[3mm] 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   340
     @{prop "False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   341
     @{text " 1. False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   342
     \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   343
   *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   344
(*<*)oops(*>*)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   345
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   346
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   347
  Another simple tactic is the function @{ML atac}, which, as shown in the previous
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   348
  section, corresponds to the assumption command.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   349
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   350
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   351
lemma shows "P \<Longrightarrow> P"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
apply(tactic {* atac 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   353
txt{*\begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   354
     @{subgoals [display]}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   355
     \end{minipage}*}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   356
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   358
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   359
  Similarly, @{ML rtac}, @{ML dtac}, @{ML etac} and @{ML ftac} correspond
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   360
  to @{text rule}, @{text drule}, @{text erule} and @{text frule}, 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   361
  respectively. Each of them takes a theorem as argument and attempts to 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   362
  apply it to a goal. Below are three self-explanatory examples.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   363
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   364
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   365
lemma shows "P \<and> Q"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
apply(tactic {* rtac @{thm conjI} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   367
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   368
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   369
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   372
lemma shows "P \<and> Q \<Longrightarrow> False"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
apply(tactic {* etac @{thm conjE} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   374
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   375
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   376
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
lemma shows "False \<and> True \<Longrightarrow> False"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
apply(tactic {* dtac @{thm conjunct2} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   381
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   382
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   383
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   387
  Note the number in each tactic call. Also as mentioned in the previous section, most 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   388
  basic tactics take such an argument; it addresses the subgoal they are analysing. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   389
  In the proof below, we first split up the conjunction in  the second subgoal by 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   390
  focusing on this subgoal first.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   391
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   392
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   393
lemma shows "Foo" and "P \<and> Q"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   394
apply(tactic {* rtac @{thm conjI} 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   395
txt {*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   396
      @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   397
      \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   398
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   399
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   400
text {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   401
  (FIXME: is it important to get the number of subgoals?)
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   402
 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   403
  The function @{ML resolve_tac} is similar to @{ML rtac}, except that it
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   404
  expects a list of theorems as arguments. From this list it will apply the
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   405
  first applicable theorem (later theorems that are also applicable can be
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   406
  explored via the lazy sequences mechanism). Given the code
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   409
ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   410
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   411
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   412
  an example for @{ML resolve_tac} is the following proof where first an outermost 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   413
  implication is analysed and then an outermost conjunction.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   414
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   415
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   416
lemma shows "C \<longrightarrow> (A \<and> B)" and "(A \<longrightarrow> B) \<and> C"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   417
apply(tactic {* resolve_tac_xmp 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   418
apply(tactic {* resolve_tac_xmp 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   419
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   420
     @{subgoals [display]} 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   421
     \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   422
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   423
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   424
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   425
  Similarly versions taking a list of theorems exist for the tactics 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   426
  @{ML dtac} (@{ML dresolve_tac}), @{ML etac} (@{ML eresolve_tac}) and so on.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   427
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   428
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   429
  Another simple tactic is @{ML cut_facts_tac}. It inserts a list of theorems
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   430
  into the assumptions of the current goal state. For example
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   431
*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   432
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   433
lemma shows "True \<noteq> False"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   434
apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   435
txt{*produces the goal state\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   436
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   437
     \begin{minipage}{\textwidth}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   438
     @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   439
     \end{minipage}*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   440
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   441
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   442
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   443
  Since rules are applied using higher-order unification, an automatic proof
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   444
  procedure might become too fragile, if it just applies inference rules as 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   445
  shown above. The reason is that a number of rules introduce meta-variables 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   446
  into the goal state. Consider for example the proof
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   447
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   448
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   449
lemma shows "\<forall>x\<in>A. P x \<Longrightarrow> Q x"
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   450
apply(tactic {* dtac @{thm bspec} 1 *})
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   451
txt{*\begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   452
     @{subgoals [display]} 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   453
     \end{minipage}*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   454
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   455
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   456
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   457
  where the application of Rule @{text bspec} generates two subgoals involving the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   458
  meta-variable @{text "?x"}. Now, if you are not careful, tactics 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   459
  applied to the first subgoal might instantiate this meta-variable in such a 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   460
  way that the second subgoal becomes unprovable. If it is clear what the @{text "?x"}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   461
  should be, then this situation can be avoided by introducing a more
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   462
  constraint version of the @{text bspec}-rule. Such constraints can be given by
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   463
  pre-instantiating theorems with other theorems. One function to do this is
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   464
  @{ML RS}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   465
  
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   466
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   467
  "@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   468
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   469
  which in the example instantiates the first premise of the @{text conjI}-rule 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   470
  with the rule @{text disjI1}. If the instantiation is impossible, as in the 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   471
  case of
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   472
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   473
  @{ML_response_fake_both [display,gray]
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   474
  "@{thm conjI} RS @{thm mp}" 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   475
"*** Exception- THM (\"RSN: no unifiers\", 1, 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   476
[\"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\", \"\<lbrakk>?P \<longrightarrow> ?Q; ?P\<rbrakk> \<Longrightarrow> ?Q\"]) raised"}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   477
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   478
  then the function raises an exception. The function @{ML RSN} is similar to @{ML RS}, but 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   479
  takes an additional number as argument that makes explicit which premise 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   480
  should be instantiated. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   481
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   482
  To improve readability of the theorems we produce below, we shall use 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   483
  the following function
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   484
*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   485
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   486
ML{*fun no_vars ctxt thm =
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   487
let 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   488
  val ((_, [thm']), _) = Variable.import_thms true [thm] ctxt
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   489
in
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   490
  thm'
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   491
end*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   493
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   494
  that transform the schematic variables of a theorem into free variables.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   495
  Using this function for the first @{ML RS}-expression above would produce
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   496
  the more readable result:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   497
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   498
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   499
  "no_vars @{context} (@{thm disjI1} RS @{thm conjI})" "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> (P \<or> Qa) \<and> Q"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   500
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   501
  If you want to instantiate more than one premise of a theorem, you can use 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   502
  the function @{ML MRS}:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   503
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   504
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   505
  "no_vars @{context} ([@{thm disjI1}, @{thm disjI2}] MRS @{thm conjI})" 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   506
  "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> (P \<or> Qa) \<and> (Pa \<or> Q)"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   507
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   508
  If you need to instantiate lists of theorems, you can use the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   509
  functions @{ML RL} and @{ML MRL}. For example in the code below,
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   510
  every theorem in the second list is instantiated with every 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   511
  theorem in the first.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   512
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   513
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   514
   "[@{thm impI}, @{thm disjI2}] RL [@{thm conjI}, @{thm disjI1}]" 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   515
"[\<lbrakk>P \<Longrightarrow> Q; Qa\<rbrakk> \<Longrightarrow> (P \<longrightarrow> Q) \<and> Qa,
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   516
 \<lbrakk>Q; Qa\<rbrakk> \<Longrightarrow> (P \<or> Q) \<and> Qa,
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   517
 (P \<Longrightarrow> Q) \<Longrightarrow> (P \<longrightarrow> Q) \<or> Qa,
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   518
 Q \<Longrightarrow> (P \<or> Q) \<or> Qa]"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   519
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   520
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   521
  The combinators for instantiating theorems are defined in @{ML_file "Pure/drule.ML"}.
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   522
  \end{readmore}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   523
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   524
  Often proofs on the ML-level involve elaborate operations on assumptions and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   525
  @{text "\<And>"}-quantified variables. To do such operations using the basic tactics 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   526
  is very unwieldy and brittle. Some convenience and
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   527
  safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   528
  and binds the various components of a goal state to a record. 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   529
  To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   530
  takes a record and just prints out the content of this record (using the 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   531
  string transformation functions from in Section~\ref{sec:printing}). Consider
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   532
  now the proof:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   533
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   534
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   535
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   536
\begin{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   537
\begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   538
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   539
ML{*fun sp_tac {prems, params, asms, concl, context, schematics} = 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   540
  let 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   541
    val str_of_params = str_of_cterms context params
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   542
    val str_of_asms = str_of_cterms context asms
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   543
    val str_of_concl = str_of_cterm context concl
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   544
    val str_of_prems = str_of_thms context prems   
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   545
    val str_of_schms = str_of_cterms context (snd schematics)    
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   546
 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   547
    val _ = (warning ("params: " ^ str_of_params);
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   548
             warning ("schematics: " ^ str_of_schms);
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   549
             warning ("assumptions: " ^ str_of_asms);
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   550
             warning ("conclusion: " ^ str_of_concl);
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   551
             warning ("premises: " ^ str_of_prems))
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   552
  in
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   553
    no_tac 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   554
  end*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   555
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   556
\end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   557
\caption{A function that prints out the various parameters provided by the tactic
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   558
 @{ML SUBPROOF}. It uses the functions defined in Section~\ref{sec:printing} for
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   559
  extracting strings from @{ML_type cterm}s and @{ML_type thm}s.\label{fig:sptac}}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   560
\end{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   561
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   562
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   563
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   564
lemma shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x"
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   565
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   566
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   567
txt {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   568
  The tactic produces the following printout:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   569
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   570
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   571
  \begin{tabular}{ll}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   572
  params:      & @{term x}, @{term y}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   573
  schematics:  & @{term z}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   574
  assumptions: & @{term "A x y"}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   575
  conclusion:  & @{term "B y x \<longrightarrow> C (z y) x"}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   576
  premises:    & @{term "A x y"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   577
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   578
  \end{quote}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   579
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   580
  Note in the actual output the brown colour of the variables @{term x} and 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   581
  @{term y}. Although they are parameters in the original goal, they are fixed inside
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   582
  the subproof. By convention these fixed variables are printed in brown colour.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   583
  Similarly the schematic variable @{term z}. The assumption, or premise, 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   584
  @{prop "A x y"} is bound as @{ML_type cterm} to the record-variable 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   585
  @{text asms}, but also as @{ML_type thm} to @{text prems}.
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   586
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   587
  Notice also that we had to append @{text [quotes] "?"} to the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   588
  \isacommand{apply}-command. The reason is that @{ML SUBPROOF} normally
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   589
  expects that the subgoal is solved completely.  Since in the function @{ML
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   590
  sp_tac} we returned the tactic @{ML no_tac}, the subproof obviously
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   591
  fails. The question-mark allows us to recover from this failure in a
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   592
  graceful manner so that the warning messages are not overwritten by an 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   593
  ``empty sequence'' error message.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   594
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   595
  If we continue the proof script by applying the @{text impI}-rule
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   596
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   597
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   598
apply(rule impI)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   599
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   600
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   601
txt {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   602
  then the tactic prints out: 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   603
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   604
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   605
  \begin{tabular}{ll}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   606
  params:      & @{term x}, @{term y}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   607
  schematics:  & @{term z}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   608
  assumptions: & @{term "A x y"}, @{term "B y x"}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   609
  conclusion:  & @{term "C (z y) x"}\\
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   610
  premises:    & @{term "A x y"}, @{term "B y x"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   611
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   612
  \end{quote}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   613
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   614
(*<*)oops(*>*)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   615
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   616
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   617
  Now also @{term "B y x"} is an assumption bound to @{text asms} and @{text prems}.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   618
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   619
  One convenience of @{ML SUBPROOF} is that we can apply the assumptions
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   620
  using the usual tactics, because the parameter @{text prems} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   621
  contains them as theorems. With this you can easily 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   622
  implement a tactic that behaves almost like @{ML atac}:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   623
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   624
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   625
ML{*val atac' = SUBPROOF (fn {prems, ...} => resolve_tac prems 1)*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   626
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   627
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   628
  If you apply @{ML atac'} to the next lemma
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   629
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   630
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   631
lemma shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   632
apply(tactic {* atac' @{context} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   633
txt{* it will produce
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   634
      @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   635
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   636
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   637
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   638
  The restriction in this tactic which is not present in @{ML atac} is 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   639
  that it cannot instantiate any
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   640
  schematic variable. This might be seen as a defect, but it is actually
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   641
  an advantage in the situations for which @{ML SUBPROOF} was designed: 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   642
  the reason is that, as mentioned before, instantiation of schematic variables can affect 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   643
  several goals and can render them unprovable. @{ML SUBPROOF} is meant 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   644
  to avoid this.
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   645
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   646
  Notice that @{ML atac'} inside @{ML SUBPROOF} calls @{ML resolve_tac} with 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   647
  the subgoal number @{text "1"} and also the outer call to @{ML SUBPROOF} in 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   648
  the \isacommand{apply}-step uses @{text "1"}. This is another advantage 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   649
  of @{ML SUBPROOF}: the addressing  inside it is completely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   650
  local to the tactic inside the subproof. It is therefore possible to also apply 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   651
  @{ML atac'} to the second goal by just writing:
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   652
*}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   653
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   654
lemma shows "True" and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   655
apply(tactic {* atac' @{context} 2 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   656
apply(rule TrueI)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   657
done
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   658
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   659
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   661
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   662
  The function @{ML SUBPROOF} is defined in @{ML_file "Pure/subgoal.ML"} and
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   663
  also described in \isccite{sec:results}. 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   664
  \end{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   665
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   666
  A similar but less powerful function than @{ML SUBPROOF} is @{ML SUBGOAL}. 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   667
  It allows you to inspect a given  subgoal. With this you can implement 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   668
  a tactic that applies a rule according to the topmost logic connective in the 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   669
  subgoal (to illustrate this we only analyse a few connectives). The code
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   670
  of the tactic is as follows.\label{tac:selecttac}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
114
13fd0a83d3c3 properly handled linenumbers in ML-text and Isar-proofs
Christian Urban <urbanc@in.tum.de>
parents: 109
diff changeset
   673
ML %linenosgray{*fun select_tac (t,i) =
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   674
  case t of
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   675
     @{term "Trueprop"} $ t' => select_tac (t',i)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   676
   | @{term "op \<Longrightarrow>"} $ _ $ t' => select_tac (t',i)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   677
   | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   678
   | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   679
   | @{term "Not"} $ _ => rtac @{thm notI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   680
   | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   681
   | _ => all_tac*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   682
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   683
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   684
  The input of the function is a term representing the subgoal and a number
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   685
  specifying the subgoal of interest. In line 3 you need to descend under the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   686
  outermost @{term "Trueprop"} in order to get to the connective you like to
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   687
  analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   688
  analysed. Similarly with meta-implications in the next line.  While for the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   689
  first five patterns we can use the @{text "@term"}-antiquotation to
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   690
  construct the patterns, the pattern in Line 8 cannot be constructed in this
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   691
  way. The reason is that an antiquotation would fix the type of the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   692
  quantified variable. So you really have to construct the pattern using the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   693
  basic term-constructors. This is not necessary in other cases, because their type
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   694
  is always fixed to function types involving only the type @{typ bool}. For the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   695
  final pattern, we chose to just return @{ML all_tac}. Consequently, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   696
  @{ML select_tac} never fails.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   697
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   698
  Let us now see how to apply this tactic. Consider the four goals:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   699
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   700
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   701
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   702
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   703
apply(tactic {* SUBGOAL select_tac 4 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   704
apply(tactic {* SUBGOAL select_tac 3 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   705
apply(tactic {* SUBGOAL select_tac 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   706
apply(tactic {* SUBGOAL select_tac 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   707
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   708
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   709
      \end{minipage} *}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   710
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   711
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   712
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   713
  where in all but the last the tactic applied an introduction rule. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   714
  Note that we applied the tactic to the goals in ``reverse'' order. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   715
  This is a trick in order to be independent from the subgoals that are 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   716
  produced by the rule. If we had applied it in the other order 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   717
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   718
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   719
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   720
apply(tactic {* SUBGOAL select_tac 1 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   721
apply(tactic {* SUBGOAL select_tac 3 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   722
apply(tactic {* SUBGOAL select_tac 4 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   723
apply(tactic {* SUBGOAL select_tac 5 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   724
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   725
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   726
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   727
  then we have to be careful to not apply the tactic to the two subgoals produced by 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   728
  the first goal. To do this can result in quite messy code. In contrast, 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   729
  the ``reverse application'' is easy to implement.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   730
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   731
  Of course, this example is contrived: there are much simpler methods available in 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   732
  Isabelle for implementing a proof procedure analysing a goal according to its topmost
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   733
  connective. These simpler methods use tactic combinators, which we will explain 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   734
  in the next section.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   735
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   736
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   737
section {* Tactic Combinators *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   738
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   739
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   740
  The purpose of tactic combinators is to build compound tactics out of
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   741
  smaller tactics. In the previous section we already used @{ML THEN}, which
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   742
  just strings together two tactics in a sequence. For example:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   745
lemma shows "(Foo \<and> Bar) \<and> False"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   746
apply(tactic {* rtac @{thm conjI} 1 THEN rtac @{thm conjI} 1 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   747
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   748
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   749
       \end{minipage} *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   750
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   751
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   752
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   753
  If you want to avoid the hard-coded subgoal addressing, then you can use
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   754
  the ``primed'' version of @{ML THEN}. For example:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   755
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   757
lemma shows "(Foo \<and> Bar) \<and> False"
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   758
apply(tactic {* (rtac @{thm conjI} THEN' rtac @{thm conjI}) 1 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   759
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   760
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   761
       \end{minipage} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   764
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   765
  Here you only have to specify the subgoal of interest only once and
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   766
  it is consistently applied to the component tactics.
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   767
  For most tactic combinators such a ``primed'' version exists and
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   768
  in what follows we will usually prefer it over the ``unprimed'' one. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   769
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   770
  If there is a list of tactics that should all be tried out in 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   771
  sequence, you can use the combinator @{ML EVERY'}. For example
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   772
  the function @{ML foo_tac'} from page~\pageref{tac:footacprime} can also 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   773
  be written as:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   774
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   775
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   776
ML{*val foo_tac'' = EVERY' [etac @{thm disjE}, rtac @{thm disjI2}, 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   777
                        atac, rtac @{thm disjI1}, atac]*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   778
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   779
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   780
  There is even another way of implementing this tactic: in automatic proof
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   781
  procedures (in contrast to tactics that might be called by the user) there
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   782
  are often long lists of tactics that are applied to the first
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   783
  subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' 1"}, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   784
  you can also just write
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   785
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   786
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   787
ML{*val foo_tac1 = EVERY1 [etac @{thm disjE}, rtac @{thm disjI2}, 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   788
                       atac, rtac @{thm disjI1}, atac]*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   789
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   790
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   791
  and call @{ML foo_tac1}.  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   792
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   793
  With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML EVERY1} it must be
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   794
  guaranteed that all component tactics successfully apply; otherwise the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   795
  whole tactic will fail. If you rather want to try out a number of tactics,
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   796
  then you can use the combinator @{ML ORELSE'} for two tactics, and @{ML
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   797
  FIRST'} (or @{ML FIRST1}) for a list of tactics. For example, the tactic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   798
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   799
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   800
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   801
ML{*val orelse_xmp = (rtac @{thm disjI1} ORELSE' rtac @{thm conjI})*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   802
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   803
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   804
  will first try out whether rule @{text disjI} applies and after that 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   805
  @{text conjI}. To see this consider the proof
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   806
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   807
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   808
lemma shows "True \<and> False" and "Foo \<or> Bar"
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   809
apply(tactic {* orelse_xmp 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   810
apply(tactic {* orelse_xmp 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   811
txt {* which results in the goal state
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   812
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   813
       \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   814
       @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   815
       \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   816
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   818
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   821
  Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac} 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   822
  as follows:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   823
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   824
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   825
ML{*val select_tac' = FIRST' [rtac @{thm conjI}, rtac @{thm impI}, 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   826
                          rtac @{thm notI}, rtac @{thm allI}, K all_tac]*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   827
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   828
text {*
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   829
  Since we like to mimic the behaviour of @{ML select_tac} as closely as possible, 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   830
  we must include @{ML all_tac} at the end of the list, otherwise the tactic will
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   831
  fail if no rule applies (we also have to wrap @{ML all_tac} using the 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   832
  @{ML K}-combinator, because it does not take a subgoal number as argument). You can 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   833
  test the tactic on the same goals:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   834
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   835
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   836
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   837
apply(tactic {* select_tac' 4 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   838
apply(tactic {* select_tac' 3 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   839
apply(tactic {* select_tac' 2 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   840
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   841
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   842
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   843
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   844
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   845
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   846
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   847
  Since such repeated applications of a tactic to the reverse order of 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   848
  \emph{all} subgoals is quite common, there is the tactic combinator 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   849
  @{ML ALLGOALS} that simplifies this. Using this combinator you can simply 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   850
  write: *}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   851
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   852
lemma shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   853
apply(tactic {* ALLGOALS select_tac' *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   854
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   855
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   856
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   857
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   858
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   859
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   860
  Remember that we chose to implement @{ML select_tac'} so that it 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   861
  always  succeeds. This can be potentially very confusing for the user, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   862
  for example,  in cases where the goal is the form
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   863
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   864
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   865
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   866
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   867
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   868
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   869
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   870
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   871
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   872
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   873
  In this case no rule applies. The problem for the user is that there is little 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   874
  chance to see whether or not progress in the proof has been made. By convention
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   875
  therefore, tactics visible to the user should either change something or fail.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   876
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   877
  To comply with this convention, we could simply delete the @{ML "K all_tac"}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   878
  from the end of the theorem list. As a result @{ML select_tac'} would only
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   879
  succeed on goals where it can make progress. But for the sake of argument,
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   880
  let us suppose that this deletion is \emph{not} an option. In such cases, you can
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   881
  use the combinator @{ML CHANGED} to make sure the subgoal has been changed
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   882
  by the tactic. Because now
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   883
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   884
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   885
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   886
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   887
apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   888
txt{* gives the error message:
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   889
  \begin{isabelle}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   890
  @{text "*** empty result sequence -- proof command failed"}\\
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   891
  @{text "*** At command \"apply\"."}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   892
  \end{isabelle} 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   893
*}(*<*)oops(*>*)
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   894
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   895
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   896
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   897
  We can further extend @{ML select_tac'} so that it not just applies to the topmost
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   898
  connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   899
  completely. For this you can use the tactic combinator @{ML REPEAT}. As an example 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   900
  suppose the following tactic
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   901
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   902
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   903
ML{*val repeat_xmp = REPEAT (CHANGED (select_tac' 1)) *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   904
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   905
text {* which applied to the proof *}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   906
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   907
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   908
apply(tactic {* repeat_xmp *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   909
txt{* produces
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   910
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   911
      \begin{minipage}{\textwidth}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   912
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   913
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   914
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   915
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   916
text {*
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   917
  Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED}, 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   918
  because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   919
  tactic as long as it succeeds). The function
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   920
  @{ML REPEAT1} is similar, but runs the tactic at least once (failing if 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   921
  this is not possible).
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   922
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   923
  If you are after the ``primed'' version of @{ML repeat_xmp} then you 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   924
  need to implement it as
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   925
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   926
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   927
ML{*val repeat_xmp' = REPEAT o CHANGED o select_tac'*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   928
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   929
text {*
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   930
  since there are no ``primed'' versions of @{ML REPEAT} and @{ML CHANGED}.
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   931
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   932
  If you look closely at the goal state above, the tactics @{ML repeat_xmp}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   933
  and @{ML repeat_xmp'} are not yet quite what we are after: the problem is
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   934
  that goals 2 and 3 are not analysed. This is because the tactic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   935
  is applied repeatedly only to the first subgoal. To analyse also all
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   936
  resulting subgoals, you can use the tactic combinator @{ML REPEAT_ALL_NEW}. 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   937
  Suppose the tactic
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   938
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   939
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   940
ML{*val repeat_all_new_xmp = REPEAT_ALL_NEW (CHANGED o select_tac')*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   941
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   942
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   943
  you see that the following goal
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   944
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   945
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   946
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   947
apply(tactic {* repeat_all_new_xmp 1 *})
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   948
txt{* \begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   949
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   950
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   951
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   952
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   953
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   954
  is completely analysed according to the theorems we chose to
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   955
  include in @{ML select_tac'}. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   956
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   957
  Recall that tactics produce a lazy sequence of successor goal states. These
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   958
  states can be explored using the command \isacommand{back}. For example
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   959
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   960
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   961
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   962
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   963
apply(tactic {* etac @{thm disjE} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   964
txt{* applies the rule to the first assumption yielding the goal state:\smallskip
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   965
      
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   966
      \begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   967
      @{subgoals [display]}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   968
      \end{minipage}\smallskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   969
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   970
      After typing
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   971
  *}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   972
(*<*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   973
oops
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   974
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   975
apply(tactic {* etac @{thm disjE} 1 *})
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   976
(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   977
back
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   978
txt{* the rule now applies to the second assumption.\smallskip
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   979
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   980
      \begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   981
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   982
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   983
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   984
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   985
text {*
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   986
  Sometimes this leads to confusing behaviour of tactics and also has
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   987
  the potential to explode the search space for tactics.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   988
  These problems can be avoided by prefixing the tactic with the tactic 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   989
  combinator @{ML DETERM}.
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   990
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   991
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   992
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   993
apply(tactic {* DETERM (etac @{thm disjE} 1) *})
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   994
txt {*\begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   995
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   996
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   997
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   998
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   999
  This combinator will prune the search space to just the first successful application.
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1000
  Attempting to apply \isacommand{back} in this goal states gives the
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1001
  error message:
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1002
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1003
  \begin{isabelle}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1004
  @{text "*** back: no alternatives"}\\
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1005
  @{text "*** At command \"back\"."}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1006
  \end{isabelle}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1007
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1008
  \begin{readmore}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1009
  Most tactic combinators described in this section are defined in @{ML_file "Pure/tctical.ML"}.
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1010
  \end{readmore}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1011
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1012
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1013
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1014
section {* Rewriting and Simplifier Tactics *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1015
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1016
text {*
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
  @{ML rewrite_goals_tac}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1018
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
  @{ML ObjectLogic.full_atomize_tac}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1020
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
  @{ML ObjectLogic.rulify_tac}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1022
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1023
  Something about simprocs.
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1024
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1027
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1028
section {* Structured Proofs *}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1029
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1030
lemma True
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1031
proof
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1032
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1033
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1034
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1035
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1036
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1037
    then have "A & B" ..
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1038
    then have C by (rule r)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1039
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1040
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1041
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1042
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1043
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1044
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1045
    note conjI [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1046
    note r [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1047
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1048
oops
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1049
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1050
ML {* fun prop ctxt s =
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1051
  Thm.cterm_of (ProofContext.theory_of ctxt) (Syntax.read_prop ctxt s) *}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1052
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1053
ML {* 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1054
  val ctxt0 = @{context};
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1055
  val ctxt = ctxt0;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1056
  val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1057
  val ([r], ctxt) = Assumption.add_assumes [prop ctxt "A & B \<Longrightarrow> C"] ctxt;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1058
  val (this, ctxt) = Assumption.add_assumes [prop ctxt "A", prop ctxt "B"] ctxt;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1059
  val this = [@{thm conjI} OF this]; 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1060
  val this = r OF this;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1061
  val this = Assumption.export false ctxt ctxt0 this 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1062
  val this = Variable.export ctxt ctxt0 [this] 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1063
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
  1066
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
end