CookBook/Tactical.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 09 Feb 2009 04:18:14 +0000
changeset 107 258ce361ba1b
parent 105 f49dc7e96235
child 108 8bea3f74889d
permissions -rw-r--r--
polished and more material in the tactic chapter
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 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  into ML. Consider the following proof.
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). 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  The functions @{ML etac}, @{ML rtac} and @{ML atac} correspond to 
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}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    65
  and the file @{ML_file "Pure/goal.ML"}. For more information about the
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    66
  internals of goals see \isccite{sec:tactical-goals}.  See @{ML_file
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    67
  "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
    68
  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
    69
  Isabelle Reference Manual.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    72
  Note that we used antiquotations for referencing the theorems. Many theorems
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    73
  also have ML-bindings with the same name. Therefore, we could also just have
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    74
  written @{ML "etac disjE 1"}, or in case there are no ML-binding obtained
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    75
  the theorem dynamically using the function @{ML thm}; for example 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    76
  @{ML "etac (thm \"disjE\") 1"}. Both ways however are considered bad style. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    77
  The reason
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    78
  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
    79
  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
    80
  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
    81
  then the theorems are fixed statically at compile-time.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
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
    83
  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
    84
  necessary to test a tactic on examples. This can be conveniently 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  Consider the following sequence of tactics
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
ML{*val foo_tac = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
      (etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
       THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
       THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
       THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
       THEN atac 1)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
text {* and the Isabelle proof: *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
lemma "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
apply(tactic {* foo_tac *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   103
  By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   104
  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
   105
  any other function that returns a tactic.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
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
   107
  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
   108
  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
   109
  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
   110
  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
   111
  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
   112
  you can write\label{tac:footacprime}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
ML{*val foo_tac' = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
      (etac @{thm disjE} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
       THEN' rtac @{thm disjI2} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
       THEN' atac 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
       THEN' rtac @{thm disjI1} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
       THEN' atac)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
text {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  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
   124
  called. So in the next proof you can first discharge the second subgoal, and
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   125
  after that the first.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
   and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
apply(tactic {* foo_tac' 2 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
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
   135
  This kind of addressing is more difficult to achieve when the goal is 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   136
  hard-coded inside the tactic. For every operator that combines tactics 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   137
  (@{ML THEN} is only one such operator) a ``primed'' version exists.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   138
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
   139
  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
   140
  analysing goals being only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   141
  of this form, then they throw the error message:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   142
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   143
  \begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   144
  @{text "*** empty result sequence -- proof command failed"}\\
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   145
  @{text "*** At command \"apply\"."}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   146
  \end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   147
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   148
  Meaning the tactics failed. The reason for this error message is that tactics 
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   149
  are functions that map a goal state to a (lazy) sequence of successor states, 
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
   150
  hence the type of a tactic is:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  @{text [display, gray] "type tactic = thm -> thm Seq.seq"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   154
  It is custom that if a tactic fails, it should return the empty sequence: 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   155
  therefore your own tactics should not raise exceptions willy-nilly. Only
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   156
  in very grave failure situations should a tactic raise the exception 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   157
  @{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
   158
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
  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
   160
  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
   161
*}
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
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
   164
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
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
   166
  which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   167
  as a single member sequence; @{ML all_tac} 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
   168
*}
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
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
   171
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
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   173
  which means it always succeeds (but also does not make any progress 
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
   174
  with the proof).
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   176
  The lazy list of possible successor states shows through at the user-level 
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   177
  of Isabelle when using the command \isacommand{back}. For instance in the 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   178
  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
   179
  @{ML foo_tac'}: either using the first assumption or the second.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
back
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
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
   187
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   189
  By using \isacommand{back}, we construct the proof that uses the
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   190
  second assumption. In more interesting situations, only by exploring 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   191
  different possibilities one might be able to find a proof.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   192
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  sequences. However in day-to-day Isabelle programming, one rarely 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  constructs sequences explicitly, but uses the predefined functions
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   197
  instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   200
  It might be surprising that tactics, which transform
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
   201
  one proof state to the next, are functions from theorems to theorem 
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
   202
  (sequences). The surprise resolves by knowing that every 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   203
  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
   204
  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
   205
  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
   206
*}
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
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
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
   209
 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
   210
   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
   211
 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
   212
   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
   213
 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
   214
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
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
   216
  which prints out the given theorem (using the string-function defined 
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
   217
  in Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. We
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   218
  are now in the position to inspect every proof state in a proof. Consider 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   219
  the proof below: on the left-hand side we show the goal state as shown 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   220
  by Isabelle, on the right-hand side the print out from @{ML my_print_tac}.
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
   221
*}
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
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   223
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
   224
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
   225
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
   226
txt{* \small 
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
   227
      \begin{tabular}{@ {}l@ {}p{0.7\textwidth}@ {}}
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
   228
      \begin{minipage}[t]{0.3\textwidth}
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
   229
      @{subgoals [display]} 
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
   230
      \end{minipage} &
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
   231
      \hfill@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
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
   232
      \end{tabular}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
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
   235
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
   236
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
   237
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
txt{* \small 
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
      \begin{tabular}{@ {}l@ {}p{0.76\textwidth}@ {}}
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
   240
      \begin{minipage}[t]{0.26\textwidth}
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]} 
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
   242
      \end{minipage} &
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
   243
      \hfill@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
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
   244
      \end{tabular}
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
   245
*}
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
   246
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
   247
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
   248
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
   249
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
   250
txt{* \small 
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
      \begin{tabular}{@ {}l@ {}p{0.7\textwidth}@ {}}
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
      \begin{minipage}[t]{0.3\textwidth}
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
      @{subgoals [display]} 
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
      \end{minipage} &
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
      \hfill@{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
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
   256
      \end{tabular}
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
*}
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
   258
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
   259
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
   260
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
   261
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
   262
txt{* \small 
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
   263
      \begin{tabular}{@ {}l@ {}p{0.7\textwidth}@ {}}
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
   264
      \begin{minipage}[t]{0.3\textwidth}
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
   265
      @{subgoals [display]} 
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
   266
      \end{minipage} &
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
      \hfill@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}
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
      \end{tabular}
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
*}
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
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
done
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
   272
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
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
   274
  As can be seen, internally every goal state is an implication of the form
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
   275
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
   276
  @{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
   277
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
   278
  where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are the  
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   279
  subgoals. So after setting up the lemma, the goal state is always of the form
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   280
  @{text "C \<Longrightarrow> (C)"}, and when the proof is finished we are left with @{text "(C)"}. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   281
  Since the goal @{term C} can potentially be an implication, 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   282
  there is a ``protector'' wrapped around it (in from of an outermost constant 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   283
  @{text "Const (\"prop\", bool \<Rightarrow> bool)"} applied to each goal;
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   284
  however this constant is invisible in the print out above). This 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   285
  prevents that premises of @{text C} are mis-interpreted as open subgoals. 
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
   286
  While tactics can operate on the subgoals (the @{text "A\<^isub>i"} above), they 
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
   287
  are expected to leave the conclusion @{term C} intact, with the 
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
   288
  exception of possibly instantiating schematic variables. 
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
   289
 
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
   290
*}
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
   291
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
   292
section {* Simple Tactics *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   294
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   295
  A simple tactic is @{ML print_tac}, which is quite useful for low-level debugging of tactics.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   296
  It just prints out a message and the current goal state.
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   297
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   298
 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   299
lemma shows "False \<Longrightarrow> True"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   300
apply(tactic {* print_tac "foo message" *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   301
txt{*\begin{minipage}{\textwidth}\small
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   302
     @{text "foo message"}\\[3mm] 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   303
     @{prop "False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   304
     @{text " 1. False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   305
     \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   306
   *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   307
(*<*)oops(*>*)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   308
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   309
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   310
  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
   311
  section, corresponds to the assumption command.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   312
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   313
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   314
lemma shows "P \<Longrightarrow> P"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
apply(tactic {* atac 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   318
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   319
  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
   320
  to @{text rule}, @{text drule}, @{text erule} and @{text frule}, 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   321
  respectively. Each of them takes a theorem as argument. Below are three 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   322
  examples with the resulting goal state. How
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   323
  they work should be self-explanatory.  
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   324
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   325
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   326
lemma shows "P \<and> Q"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
apply(tactic {* rtac @{thm conjI} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   328
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   329
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   330
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   333
lemma shows "P \<and> Q \<Longrightarrow> False"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
apply(tactic {* etac @{thm conjE} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   335
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   336
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   337
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
lemma shows "False \<and> True \<Longrightarrow> False"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
apply(tactic {* dtac @{thm conjunct2} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   342
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   343
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   344
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   348
  As mentioned in the previous section, most basic tactics take a number as 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   349
  argument, which addresses the subgoal they are analysing. In the proof below,
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   350
  we first analyse the second subgoal by focusing on this subgoal first.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   351
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   352
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   353
lemma shows "Foo" and "P \<and> Q"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   354
apply(tactic {* rtac @{thm conjI} 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   355
txt {*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   356
      @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   357
      \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   358
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   359
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   360
text {* 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   361
  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
   362
  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
   363
  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
   364
  explored via the lazy sequences mechanism). Given the code
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   367
ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   368
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   369
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   370
  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
   371
  implication is analysed and then an outermost conjunction.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   372
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   373
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   374
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
   375
apply(tactic {* resolve_tac_xmp 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   376
apply(tactic {* resolve_tac_xmp 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   377
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   378
     @{subgoals [display]} 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   379
     \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   380
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   381
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   382
text {* 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   383
  Similarly versions exists for @{ML atac} (@{ML assume_tac}), @{ML etac} 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   384
  (@{ML eresolve_tac}) and so on. 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   385
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   386
  Another simple tactic is @{ML cut_facts_tac}. It inserts a list of theorems
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   387
  into the assumptions of the current goal state. For example:
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   388
*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   389
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   390
lemma shows "True \<noteq> False"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   391
apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   392
txt{*\begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   393
     @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   394
     \end{minipage}*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   395
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   396
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   397
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   398
  Since rules are applied using higher-order unification, an automatic proof
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   399
  procedure might become too fragile, if it just applies inference rules shown
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   400
  in the fashion above.  More constraints can be introduced by
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   401
  pre-instantiating theorems with other theorems. You can do this using the
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   402
  function @{ML RS}. For example
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   403
  
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   404
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   405
  "@{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
   406
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   407
  instantiates the first premise of the @{text conjI}-rule with the
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   408
  rule @{text disjI1}. If this is impossible, as in the case of
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   409
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   410
  @{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
   411
  "@{thm conjI} RS @{thm mp}" 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   412
"*** 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
   413
[\"\<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
   414
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   415
  the function raises an exception. The function @{ML RSN} is similar, but 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   416
  takes a number as argument and thus you can make explicit which premise 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   417
  should be instantiated. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   418
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   419
  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
   420
  the following function
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   421
*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   422
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   423
ML{*fun no_vars ctxt thm =
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   424
let 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   425
  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
   426
in
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   427
  thm'
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   428
end*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   430
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   431
  that transform the schematic variables of a theorem into free variables. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   432
  This means for the first @{ML RS}-expression above:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   433
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   434
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   435
  "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
   436
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   437
  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
   438
  the function @{ML MRS}:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   439
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   440
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   441
  "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
   442
  "\<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
   443
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   444
  If you need to instantiate lists of theorems, you can use the
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   445
  functions @{ML RL} and @{ML MRL}. For example in the code below
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   446
  every theorem in the first list is instantiated against every 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   447
  theorem in the second.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   448
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   449
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   450
   "[@{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
   451
"[\<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
   452
 \<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
   453
 (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
   454
 Q \<Longrightarrow> (P \<or> Q) \<or> Qa]"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   455
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   456
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   457
  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
   458
  \end{readmore}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   459
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   460
  Often proofs involve elaborate operations on assumptions and 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   461
  @{text "\<And>"}-quantified variables. To do such operations on the ML-level 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   462
  using the basic tactics is very unwieldy and brittle. Some convenience and
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   463
  safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   464
  and binds the various components of a proof state into a record. 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   465
  To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   466
  takes a record as argument and just prints out the content of this record (using the 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   467
  string transformation functions defined in Section~\ref{sec:printing}). Consider
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   468
  now the proof
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   469
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   470
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   471
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   472
\begin{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   473
\begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   474
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   475
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
   476
  let 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   477
    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
   478
    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
   479
    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
   480
    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
   481
    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
   482
 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   483
    val _ = (warning ("params: " ^ str_of_params);
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   484
             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
   485
             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
   486
             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
   487
             warning ("premises: " ^ str_of_prems))
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   488
  in
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   489
    no_tac 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   490
  end*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   491
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   492
\end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   493
\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
   494
 @{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
   495
  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
   496
\end{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   497
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   498
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   499
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   500
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
   501
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   502
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   503
txt {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   504
  which gives the printout:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   505
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   506
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   507
  \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
   508
  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
   509
  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
   510
  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
   511
  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
   512
  premises:    & @{term "A x y"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   513
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   514
  \end{quote}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   515
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   516
  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
   517
  @{term y}. Although they are parameters in the original goal, they are fixed inside
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   518
  the subproof. Similarly the schematic variable @{term z}. The assumption 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   519
  @{prop "A x y"} is bound as @{ML_type cterm} to the record-variable 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   520
  @{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
   521
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   522
  Notice also that we had to append @{text "?"} to \isacommand{apply}. The 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   523
  reason is that @{ML SUBPROOF} normally expects that the subgoal is solved completely.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   524
  Since in the function @{ML sp_tac} we returned the tactic @{ML no_tac}, the subproof
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   525
  obviously fails. The question-mark allows us to recover from this failure
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   526
  in a graceful manner so that the warning messages are not overwritten
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   527
  by an error message.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   528
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   529
  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
   530
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   531
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   532
apply(rule impI)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   533
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   534
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   535
txt {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   536
  then @{ML sp_tac} prints out 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   537
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   538
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   539
  \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
   540
  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
   541
  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
   542
  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
   543
  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
   544
  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
   545
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   546
  \end{quote}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   547
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   548
(*<*)oops(*>*)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   549
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   550
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   551
  where we now also have @{term "B y x"} as an assumption.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   552
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   553
  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
   554
  using the usual tactics, because the parameter @{text prems} 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   555
  contains them as theorems. With this we can easily 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   556
  implement a tactic that almost behaves like @{ML atac}, namely:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   557
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   558
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   559
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
   560
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   561
text {*
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   562
  If we apply it to the next lemma
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   563
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   564
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   565
lemma shows "\<And>x y. \<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
   566
apply(tactic {* atac' @{context} 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   567
txt{* we get
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   568
      @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   569
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   570
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   571
text {*
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   572
  The restriction in this tactic is 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
   573
  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
   574
  an advantage in the situations for which @{ML SUBPROOF} was designed: 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   575
  the reason is that instantiation of schematic variables can affect 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   576
  several goals and can render them unprovable. @{ML SUBPROOF} is meant 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   577
  to avoid this.
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   578
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   579
  Notice that @{ML atac'} calls @{ML resolve_tac} with the subgoal
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   580
  number @{text "1"} and also the ``outer'' call to @{ML SUBPROOF} in 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   581
  the \isacommand{apply}-step uses @{text "1"}. Another advantage 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   582
  of @{ML SUBGOAL} is that the addressing  inside it is completely 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   583
  local to the subproof inside. 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
   584
  @{ML atac'} to the second goal by just writing:
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   585
*}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   586
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   587
lemma shows "True" and "\<And>x y. \<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   588
apply(tactic {* atac' @{context} 2 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   589
apply(rule TrueI)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   590
done
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   591
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   592
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   594
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   595
  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
   596
  also described in \isccite{sec:results}. 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   597
  \end{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   598
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   599
  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
   600
  It allows you to inspect a given  subgoal. With this you can implement 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   601
  a tactic that applies a rule according to the topmost connective in the 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   602
  subgoal (to illustrate this we only analyse a few connectives). The code
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   603
  of this tactic is as follows.\label{tac:selecttac}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   606
ML %linenumbers{*fun select_tac (t,i) =
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   607
  case t of
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   608
     @{term "Trueprop"} $ t' => select_tac (t',i)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   609
   | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   610
   | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   611
   | @{term "Not"} $ _ => rtac @{thm notI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   612
   | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   613
   | _ => all_tac*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   614
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   615
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   616
  In line 3 you need to decend under the outermost @{term "Trueprop"} in order
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   617
  to get to the connective you like to analyse. Otherwise goals like @{prop "A \<and> B"} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   618
  are not properly analysed. While for the first four pattern we can use the 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   619
  @{text "@term"}-antiquotation, the pattern in Line 7 cannot be constructed
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   620
  in this way. The reason is that an antiquotation would fix the type of the 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   621
  quantified variable. So you really have to construct the pattern
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   622
  using the term-constructors. This is not necessary in  other cases, because 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   623
  their type is always something involving @{typ bool}. The final patter is 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   624
  for the case where the goal does not fall into any of the categorories before.
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   625
  In this case we chose to just return @{ML all_tac} (i.e., @{ML select_tac} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   626
  never fails). 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   627
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   628
  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
   629
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   630
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   631
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   632
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
   633
apply(tactic {* SUBGOAL select_tac 4 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   634
apply(tactic {* SUBGOAL select_tac 3 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   635
apply(tactic {* SUBGOAL select_tac 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   636
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
   637
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   638
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   639
      \end{minipage} *}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   640
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   641
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   642
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   643
  where in all but the last the tactic applied an introduction rule. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   644
  Note that we applied the tactic to subgoals in ``reverse'' order. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   645
  This is a trick in order to be independent from what subgoals are 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   646
  that are produced by the rule. If we had it applied in the other order 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   647
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   648
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   649
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
   650
apply(tactic {* SUBGOAL select_tac 1 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   651
apply(tactic {* SUBGOAL select_tac 3 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   652
apply(tactic {* SUBGOAL select_tac 4 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   653
apply(tactic {* SUBGOAL select_tac 5 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   654
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   655
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   656
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   657
  then we have to be careful to not apply the tactic to the two subgoals the
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   658
  first goal produced. To do this can result in quite messy code. In contrast, 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   659
  the ``reverse application'' is easy to implement.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   660
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   661
  However, this example is very contrived: there are much simpler methods to implement
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   662
  such a proof procedure analying a goal according to its topmost
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   663
  connective. These simpler methods use tactic combinators which will be explained 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   664
  in the next section.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   665
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   666
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   667
section {* Tactic Combinators *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   668
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   669
text {* 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   670
  The purpose of tactic combinators is to build powerful tactics out of
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   671
  smaller components. In the previous section we already used @{ML THEN} which
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   672
  strings two tactics together in sequence. For example:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   675
lemma shows "(Foo \<and> Bar) \<and> False"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   676
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
   677
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   678
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   679
       \end{minipage} *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   680
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   681
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   682
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   683
  If you want to avoid the hard-coded subgoal addressing, then you can use
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   684
  the ``primed'' version of @{ML THEN}, namely @{ML THEN'}. For example:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   685
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   687
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
   688
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
   689
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   690
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   691
       \end{minipage} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   694
text {* 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   695
  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
   696
  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
   697
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   698
  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
   699
  sequence, you can use the combinator @{ML EVERY'}. For example
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   700
  the function @{ML foo_tac'} from page~\ref{tac:footacprime} can also 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   701
  be written as
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   702
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   703
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   704
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
   705
                        atac, rtac @{thm disjI1}, atac]*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   706
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   707
text {*
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   708
  There is even another variant for @{ML foo_tac''}: in automatic proof
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   709
  procedures (in contrast to tactics that might be called by the user) 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   710
  there are often long lists of tactics that are applied to the first
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   711
  subgoal. Instead of writing the code above and then calling 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   712
  @{ML "foo_tac'' 1"}, you can also just write
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   713
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   714
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   715
ML{*val foo_tac1 = EVERY1 [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
   716
                         atac, rtac @{thm disjI1}, atac]*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   717
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   718
text {*
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   719
  With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML EVERY1} it must be 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   720
  guaranteed that  all component tactics sucessfully apply; otherwise the 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   721
  whole tactic will fail. If you rather want to try out a number of tactics, 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   722
  then you can use the combinator @{ML ORELSE'} for two tactics and @{ML FIRST'} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   723
  (or @{ML FIRST1}) for a list of tactics. For example, the tactic
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   724
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   725
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   726
ML{*val orelse_xmp = (rtac @{thm disjI1} ORELSE' rtac @{thm conjI})*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   727
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   728
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   729
  will first try out whether rule @{text disjI} applies and after that 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   730
  whether @{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
   731
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   732
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   733
lemma shows "True \<and> False" and "Foo \<or> Bar"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   734
apply(tactic {* orelse_xmp 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   735
apply(tactic {* orelse_xmp 3 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   736
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
   737
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   738
       \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   739
       @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   740
       \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   741
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   746
  Using @{ML FIRST'} we can write our @{ML select_tac} from Page~\ref{tac:selecttac} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   747
  simply as follows:
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   748
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   749
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   750
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
   751
                          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
   752
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   753
text {*
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   754
  Since we like to mimic the bahaviour of @{ML select_tac}, we must include
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   755
  @{ML all_tac} at the end (@{ML all_tac} must also be ``wrapped up'' using
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   756
  the @{ML K}-combinator as it does not take a subgoal number as argument). 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   757
  We can test the tactic on the same proof:
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   758
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   759
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   760
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
   761
apply(tactic {* select_tac' 4 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   762
apply(tactic {* select_tac' 3 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   763
apply(tactic {* select_tac' 2 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   764
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   765
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   766
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   767
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   768
(*<*)oops(*>*)
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
text {* 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   771
  and obtain the same subgoals. Since this repeated application of a tactic
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   772
  to the reverse order of \emph{all} subgoals is quite common, there is 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   773
  the tactics combinator @{ML ALLGOALS} that simplifies this. Using this 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   774
  combinator we can simply write: *}
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
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
   777
apply(tactic {* ALLGOALS select_tac' *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   778
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   779
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   780
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   781
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   782
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   783
text {*
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   784
  We chose to write @{ML select_tac'} in such a way that it always succeeds.
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   785
  This can be potetially very  confusing for the user in cases of goals 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   786
  of the form
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   787
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   788
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   789
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   790
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   791
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   792
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   793
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   794
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   795
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   796
text {*
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   797
  where no rule applies. The reason is that the user has little chance 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   798
  to see whether progress in the proof has been made or not. We could simply
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   799
  remove @{ML "K all_tac"} from the end of the list. Then the tactic would
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   800
  only apply in cases where it can make some progress. But for the sake of
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   801
  argument, let us assume that this is not an option. In such cases, you 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   802
  can use the combinator @{ML CHANGED} to make sure the subgoal has been
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   803
  changed by the tactic. Because now
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   804
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   805
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   806
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   807
apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   808
txt{* results in the usual error message for empty result sequences. *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   809
(*<*)oops(*>*)
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   810
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   811
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   812
text {*
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   813
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   814
  @{ML REPEAT} @{ML DETERM}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   816
  @{ML CHANGED}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   817
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   818
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   819
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   820
section {* Rewriting and Simplifier Tactics *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   821
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   822
text {*
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
  @{ML rewrite_goals_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
  @{ML ObjectLogic.full_atomize_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
  @{ML ObjectLogic.rulify_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   828
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   829
section {* Structured Proofs *}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   830
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   831
lemma True
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   832
proof
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   833
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   834
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   835
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   836
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   837
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   838
    then have "A & B" ..
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   839
    then have C by (rule r)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   840
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   841
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   842
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   843
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   844
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   845
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   846
    note conjI [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   847
    note r [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   848
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   849
oops
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   850
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   851
ML {* fun prop ctxt s =
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   852
  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
   853
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   854
ML {* 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   855
  val ctxt0 = @{context};
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   856
  val ctxt = ctxt0;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   857
  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
   858
  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
   859
  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
   860
  val this = [@{thm conjI} OF this]; 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   861
  val this = r OF this;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   862
  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
   863
  val this = Variable.export ctxt ctxt0 [this] 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   864
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   865
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   866
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
   867
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   868
end