CookBook/Tactical.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 25 Feb 2009 11:28:41 +0000
changeset 144 eaba1442c516
parent 142 c06885c36575
child 145 f1ba430a5e7d
permissions -rw-r--r--
new pdf file
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
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
     3
uses "infix_conv.ML"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
chapter {* Tactical Reasoning\label{chp:tactical} *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
     9
  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
    10
  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
    11
  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
    12
  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
    13
  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
    14
  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
    15
  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
    16
  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
    17
  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
    18
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
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
    21
section {* Basics of Reasoning with Tactics*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
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
    24
  To see how tactics work, let us first transcribe a simple \isacommand{apply}-style proof 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
    25
  into ML. Suppose the following proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
apply(erule disjE)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
apply(rule disjI2)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
apply(rule disjI1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  This proof translates to the following ML-code.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
@{ML_response_fake [display,gray]
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
"let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  Goal.prove ctxt [\"P\", \"Q\"] [] goal 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
   (fn _ => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
      etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
      THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
      THEN atac 1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  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
    54
  tac"} sets up a goal state for proving the goal @{text C} 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    55
  (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
    56
  assumptions @{text As} (happens to be empty) with the variables
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  @{text xs} that will be generalised once the goal is proved (in our case
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  @{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  it can make use of the local assumptions (there are none in this example). 
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    60
  The functions @{ML etac}, @{ML rtac} and @{ML atac} in the code above correspond to 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  @{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
    62
  The operator @{ML THEN} strings the tactics together. 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  \begin{readmore}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    65
  To learn more about the function @{ML Goal.prove} see \isccite{sec:results}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
    66
  and the file @{ML_file "Pure/goal.ML"}.  See @{ML_file
99
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
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    72
  Note that in the code above we use antiquotations for referencing the theorems. Many theorems
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    73
  also have ML-bindings with the same name. Therefore, we could also just have
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
    74
  written @{ML "etac disjE 1"}, or in case where there are no ML-binding obtain
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    75
  the theorem dynamically using the function @{ML thm}; for example 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
    76
  \mbox{@{ML  "etac (thm \"disjE\") 1"}}. Both ways however are considered bad style! 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    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
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   125
  subsequently 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 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   136
  hard-coded inside the tactic. For most operators that combine tactics 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   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
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   141
  of this form, then they return 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
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   148
  This means the tactics failed. The reason for this error message is that tactics 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   149
  are functions mapping a goal state to a (lazy) sequence of successor states. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   150
  Hence the type of a tactic is:
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   151
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   153
ML{*type tactic = thm -> thm Seq.seq*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   155
text {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   156
  By convention, if a tactic fails, then it should return the empty sequence. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   157
  Therefore, if you write your own tactics, they  should not raise exceptions 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   158
  willy-nilly; only in very grave failure situations should a tactic raise the 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   159
  exception @{text THM}.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   160
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   161
  The 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
   162
  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
   163
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   164
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
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
   166
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   167
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
   168
  which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   169
  up in a single member sequence; it is defined as
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   170
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   171
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
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
   173
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
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   175
  which means @{ML all_tac} always succeeds, but also does not make any progress 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   176
  with the proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   178
  The lazy list of possible successor goal states shows through at the user-level 
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   179
  of Isabelle when using the command \isacommand{back}. For instance in the 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   180
  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
   181
  @{ML foo_tac'}: either using the first assumption or the second.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
back
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
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
   189
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   191
  By using \isacommand{back}, we construct the proof that uses the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   192
  second assumption. While in the proof above, it does not really matter which 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   193
  assumption is used, in more interesting cases provability might depend
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   194
  on exploring different possibilities.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   195
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
  See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   198
  sequences. In day-to-day Isabelle programming, however, one rarely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   199
  constructs sequences explicitly, but uses the predefined tactics and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   200
  tactic combinators instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   203
  It might be surprising that tactics, which transform
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   204
  one goal state to the next, are functions from theorems to theorem 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   205
  (sequences). The surprise resolves by knowing that every 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   206
  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
   207
  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
   208
  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
   209
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   210
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
ML{*fun my_print_tac ctxt thm =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   212
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   213
  val _ = warning (str_of_thm ctxt thm)
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   214
in 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   215
  Seq.single thm
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   216
end*}
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
   217
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   218
text_raw {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   219
  \begin{figure}[p]
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   220
  \begin{isabelle}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   221
*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   222
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
   223
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
   224
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   225
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   226
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   227
      \end{minipage}\medskip      
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   228
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   229
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   230
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   231
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   232
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   233
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   234
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   235
      \end{minipage}\medskip
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
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
   238
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
   239
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
   240
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   241
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   242
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   243
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   244
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   245
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   246
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   247
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   248
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   249
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   250
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   251
      \end{minipage}\medskip
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   252
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   253
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   254
apply(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
   255
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
   256
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   257
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   258
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   259
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   260
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   261
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   262
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   263
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   264
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   265
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   266
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   267
      \end{minipage}\medskip
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   268
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   269
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   270
apply(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
   271
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
   272
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   273
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   274
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   275
      \end{minipage}\medskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   276
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   277
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   278
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   279
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   280
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   281
      @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   282
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   283
      \end{minipage}\medskip   
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   284
   *}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   285
done
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   286
text_raw {*  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   287
  \end{isabelle}
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   288
  \caption{The figure shows a proof where each intermediate goal state is
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   289
  printed by the Isabelle system and by @{ML my_print_tac}. The latter shows
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   290
  the goal state as represented internally (highlighted boxes). This
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   291
  illustrates that every goal state in Isabelle is represented by a theorem:
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   292
  when we start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}} the theorem is
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   293
  @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when we finish the proof the
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   294
  theorem is @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}.\label{fig:goalstates}}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   295
  \end{figure}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   296
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   297
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   298
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
   299
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   300
  which prints out the given theorem (using the string-function defined in
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   301
  Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   302
  tactic we are in the position to inspect every goal state in a
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   303
  proof. Consider now the proof in Figure~\ref{fig:goalstates}: as can be seen, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   304
  internally every goal state is an implication of the form
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   305
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   306
  @{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
   307
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   308
  where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   309
  the subgoals. So after setting up the lemma, the goal state is always of the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   310
  form @{text "C \<Longrightarrow> (C)"}; when the proof is finished we are left with @{text
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   311
  "(C)"}. Since the goal @{term C} can potentially be an implication, there is
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   312
  a ``protector'' wrapped around it (in from of an outermost constant @{text
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   313
  "Const (\"prop\", bool \<Rightarrow> bool)"}; however this constant
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   314
  is invisible in the figure). This prevents that premises of @{text C} are
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   315
  mis-interpreted as open subgoals. While tactics can operate on the subgoals
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   316
  (the @{text "A\<^isub>i"} above), they are expected to leave the conclusion
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   317
  @{term C} intact, with the exception of possibly instantiating schematic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   318
  variables. If you use the predefined tactics, which we describe in the next
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   319
  section, this will always be the case.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   320
 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   321
  \begin{readmore}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   322
  For more information about the internals of goals see \isccite{sec:tactical-goals}.
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   323
  \end{readmore}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   324
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
   325
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   326
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
   327
section {* Simple Tactics *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   329
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   330
  Let us start with the tactic @{ML print_tac}, which is quite useful for low-level 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   331
  debugging of tactics. It just prints out a message and the current goal state. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   332
  Processing the proof
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   333
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   334
 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   335
lemma shows "False \<Longrightarrow> True"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   336
apply(tactic {* print_tac "foo message" *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   337
txt{*gives:\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   338
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   339
     \begin{minipage}{\textwidth}\small
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   340
     @{text "foo message"}\\[3mm] 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   341
     @{prop "False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   342
     @{text " 1. False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   343
     \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   344
   *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   345
(*<*)oops(*>*)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   346
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   347
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   348
  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
   349
  section, corresponds to the assumption command.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   350
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   351
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   352
lemma shows "P \<Longrightarrow> P"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
apply(tactic {* atac 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   354
txt{*\begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   355
     @{subgoals [display]}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   356
     \end{minipage}*}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   357
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   359
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   360
  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
   361
  to @{text rule}, @{text drule}, @{text erule} and @{text frule}, 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   362
  respectively. Each of them takes a theorem as argument and attempts to 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   363
  apply it to a goal. Below are three self-explanatory examples.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   364
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   365
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   366
lemma shows "P \<and> Q"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
apply(tactic {* rtac @{thm conjI} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   368
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   369
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   370
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   373
lemma shows "P \<and> Q \<Longrightarrow> False"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
apply(tactic {* etac @{thm conjE} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   375
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   376
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   377
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
lemma shows "False \<and> True \<Longrightarrow> False"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
apply(tactic {* dtac @{thm conjunct2} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   382
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   383
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   384
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   388
  Note the number in each tactic call. Also as mentioned in the previous section, most 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   389
  basic tactics take such an argument; it addresses the subgoal they are analysing. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   390
  In the proof below, we first split up the conjunction in  the second subgoal by 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   391
  focusing on this subgoal first.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   392
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   393
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   394
lemma shows "Foo" and "P \<and> Q"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   395
apply(tactic {* rtac @{thm conjI} 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   396
txt {*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   397
      @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   398
      \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   399
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   400
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   401
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   402
  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
   403
  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
   404
  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
   405
  explored via the lazy sequences mechanism). Given the code
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   408
ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   409
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   410
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   411
  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
   412
  implication is analysed and then an outermost conjunction.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   413
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   414
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   415
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
   416
apply(tactic {* resolve_tac_xmp 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   417
apply(tactic {* resolve_tac_xmp 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   418
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   419
     @{subgoals [display]} 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   420
     \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   421
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   422
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   423
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   424
  Similarly versions taking a list of theorems exist for the tactics 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   425
  @{ML dtac} (@{ML dresolve_tac}), @{ML etac} (@{ML eresolve_tac}) and so on.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   426
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   427
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   428
  Another simple tactic is @{ML cut_facts_tac}. It inserts a list of theorems
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   429
  into the assumptions of the current goal state. For example
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   430
*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   431
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   432
lemma shows "True \<noteq> False"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   433
apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   434
txt{*produces the goal state\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   435
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   436
     \begin{minipage}{\textwidth}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   437
     @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   438
     \end{minipage}*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   439
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   440
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   441
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   442
  Since rules are applied using higher-order unification, an automatic proof
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   443
  procedure might become too fragile, if it just applies inference rules as 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   444
  shown above. The reason is that a number of rules introduce meta-variables 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   445
  into the goal state. Consider for example the proof
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   446
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   447
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   448
lemma shows "\<forall>x\<in>A. P x \<Longrightarrow> Q x"
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   449
apply(tactic {* dtac @{thm bspec} 1 *})
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   450
txt{*\begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   451
     @{subgoals [display]} 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   452
     \end{minipage}*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   453
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   454
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   455
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   456
  where the application of Rule @{text bspec} generates two subgoals involving the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   457
  meta-variable @{text "?x"}. Now, if you are not careful, tactics 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   458
  applied to the first subgoal might instantiate this meta-variable in such a 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   459
  way that the second subgoal becomes unprovable. If it is clear what the @{text "?x"}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   460
  should be, then this situation can be avoided by introducing a more
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   461
  constraint version of the @{text bspec}-rule. Such constraints can be given by
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   462
  pre-instantiating theorems with other theorems. One function to do this is
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   463
  @{ML RS}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   464
  
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   465
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   466
  "@{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
   467
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   468
  which in the example instantiates the first premise of the @{text conjI}-rule 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   469
  with the rule @{text disjI1}. If the instantiation is impossible, as in the 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   470
  case of
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   471
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   472
  @{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
   473
  "@{thm conjI} RS @{thm mp}" 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   474
"*** 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
   475
[\"\<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
   476
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   477
  then the function raises an exception. The function @{ML RSN} is similar to @{ML RS}, but 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   478
  takes an additional number as argument that makes explicit which premise 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   479
  should be instantiated. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   480
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   481
  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
   482
  the following function
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   483
*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   484
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   485
ML{*fun no_vars ctxt thm =
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   486
let 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   487
  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
   488
in
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   489
  thm'
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   490
end*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   492
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   493
  that transform the schematic variables of a theorem into free variables.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   494
  Using this function for the first @{ML RS}-expression above would produce
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   495
  the more readable result:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   496
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   497
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   498
  "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
   499
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   500
  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
   501
  the function @{ML MRS}:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   502
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   503
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   504
  "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
   505
  "\<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
   506
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   507
  If you need to instantiate lists of theorems, you can use the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   508
  functions @{ML RL} and @{ML MRL}. For example in the code below,
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   509
  every theorem in the second list is instantiated with every 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   510
  theorem in the first.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   511
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   512
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   513
   "[@{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
   514
"[\<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
   515
 \<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
   516
 (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
   517
 Q \<Longrightarrow> (P \<or> Q) \<or> Qa]"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   518
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   519
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   520
  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
   521
  \end{readmore}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   522
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   523
  Often proofs on the ML-level involve elaborate operations on assumptions and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   524
  @{text "\<And>"}-quantified variables. To do such operations using the basic tactics 
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   525
  shown so far is very unwieldy and brittle. Some convenience and
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   526
  safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   527
  and binds the various components of a goal state to a record. 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   528
  To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   529
  takes a record and just prints out the content of this record (using the 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   530
  string transformation functions from in Section~\ref{sec:printing}). Consider
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   531
  now the proof:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   532
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   533
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   534
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   535
\begin{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   536
\begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   537
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   538
ML{*fun sp_tac {prems, params, asms, concl, context, schematics} = 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   539
let 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   540
  val str_of_params = str_of_cterms context params
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   541
  val str_of_asms = str_of_cterms context asms
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   542
  val str_of_concl = str_of_cterm context concl
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   543
  val str_of_prems = str_of_thms context prems   
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   544
  val str_of_schms = str_of_cterms context (snd schematics)    
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   545
 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   546
  val _ = (warning ("params: " ^ str_of_params);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   547
           warning ("schematics: " ^ str_of_schms);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   548
           warning ("assumptions: " ^ str_of_asms);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   549
           warning ("conclusion: " ^ str_of_concl);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   550
           warning ("premises: " ^ str_of_prems))
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   551
in
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   552
  no_tac 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   553
end*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   554
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   555
\end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   556
\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
   557
 @{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
   558
  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
   559
\end{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   560
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   561
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   562
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   563
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
   564
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   565
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   566
txt {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   567
  The tactic produces the following printout:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   568
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   569
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   570
  \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
   571
  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
   572
  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
   573
  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
   574
  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
   575
  premises:    & @{term "A x y"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   576
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   577
  \end{quote}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   578
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   579
  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
   580
  @{term y}. Although they are parameters in the original goal, they are fixed inside
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   581
  the subproof. By convention these fixed variables are printed in brown colour.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   582
  Similarly the schematic variable @{term z}. The assumption, or premise, 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   583
  @{prop "A x y"} is bound as @{ML_type cterm} to the record-variable 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   584
  @{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
   585
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   586
  Notice also that we had to append @{text [quotes] "?"} to the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   587
  \isacommand{apply}-command. The reason is that @{ML SUBPROOF} normally
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   588
  expects that the subgoal is solved completely.  Since in the function @{ML
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   589
  sp_tac} we returned the tactic @{ML no_tac}, the subproof obviously
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   590
  fails. The question-mark allows us to recover from this failure in a
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   591
  graceful manner so that the warning messages are not overwritten by an 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   592
  ``empty sequence'' error message.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   593
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   594
  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
   595
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   596
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   597
apply(rule impI)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   598
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   599
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   600
txt {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   601
  then the tactic prints out: 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   602
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   603
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   604
  \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
   605
  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
   606
  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
   607
  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
   608
  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
   609
  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
   610
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   611
  \end{quote}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   612
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   613
(*<*)oops(*>*)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   614
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   615
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   616
  Now also @{term "B y x"} is an assumption bound to @{text asms} and @{text prems}.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   617
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   618
  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
   619
  using the usual tactics, because the parameter @{text prems} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   620
  contains them as theorems. With this you can easily 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   621
  implement a tactic that behaves almost like @{ML atac}:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   622
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   623
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   624
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
   625
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   626
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   627
  If you apply @{ML atac'} to the next lemma
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   628
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   629
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   630
lemma shows "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   631
apply(tactic {* atac' @{context} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   632
txt{* it will produce
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   633
      @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   634
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   635
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   636
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   637
  The restriction in this tactic which is not present in @{ML atac} is 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   638
  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
   639
  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
   640
  an advantage in the situations for which @{ML SUBPROOF} was designed: 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   641
  the reason is that, as mentioned before, instantiation of schematic variables can affect 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   642
  several goals and can render them unprovable. @{ML SUBPROOF} is meant 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   643
  to avoid this.
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   644
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   645
  Notice that @{ML atac'} inside @{ML SUBPROOF} calls @{ML resolve_tac} with 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   646
  the subgoal number @{text "1"} and also the outer call to @{ML SUBPROOF} in 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   647
  the \isacommand{apply}-step uses @{text "1"}. This is another advantage 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   648
  of @{ML SUBPROOF}: the addressing  inside it is completely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   649
  local to the tactic inside the subproof. It is therefore possible to also apply 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   650
  @{ML atac'} to the second goal by just writing:
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   651
*}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   652
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   653
lemma shows "True" and "\<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   654
apply(tactic {* atac' @{context} 2 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   655
apply(rule TrueI)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   656
done
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   657
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   658
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   660
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   661
  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
   662
  also described in \isccite{sec:results}. 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   663
  \end{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   664
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   665
  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
   666
  It allows you to inspect a given  subgoal. With this you can implement 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   667
  a tactic that applies a rule according to the topmost logic connective in the 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   668
  subgoal (to illustrate this we only analyse a few connectives). The code
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   669
  of the tactic is as follows.\label{tac:selecttac}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   670
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
114
13fd0a83d3c3 properly handled linenumbers in ML-text and Isar-proofs
Christian Urban <urbanc@in.tum.de>
parents: 109
diff changeset
   672
ML %linenosgray{*fun select_tac (t,i) =
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   673
  case t of
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   674
     @{term "Trueprop"} $ t' => select_tac (t',i)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   675
   | @{term "op \<Longrightarrow>"} $ _ $ t' => select_tac (t',i)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   676
   | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   677
   | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   678
   | @{term "Not"} $ _ => rtac @{thm notI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   679
   | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   680
   | _ => all_tac*}
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 {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   683
  The input of the function is a term representing the subgoal and a number
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   684
  specifying the subgoal of interest. In line 3 you need to descend under the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   685
  outermost @{term "Trueprop"} in order to get to the connective you like to
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   686
  analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   687
  analysed. Similarly with meta-implications in the next line.  While for the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   688
  first five patterns we can use the @{text "@term"}-antiquotation to
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   689
  construct the patterns, the pattern in Line 8 cannot be constructed in this
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   690
  way. The reason is that an antiquotation would fix the type of the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   691
  quantified variable. So you really have to construct the pattern using the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   692
  basic term-constructors. This is not necessary in other cases, because their type
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   693
  is always fixed to function types involving only the type @{typ bool}. For the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   694
  final pattern, we chose to just return @{ML all_tac}. Consequently, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   695
  @{ML select_tac} never fails.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   696
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   697
  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
   698
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   699
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   700
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   701
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
   702
apply(tactic {* SUBGOAL select_tac 4 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   703
apply(tactic {* SUBGOAL select_tac 3 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   704
apply(tactic {* SUBGOAL select_tac 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   705
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
   706
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   707
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   708
      \end{minipage} *}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   709
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   710
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   711
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   712
  where in all but the last the tactic applied an introduction rule. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   713
  Note that we applied the tactic to the goals in ``reverse'' order. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   714
  This is a trick in order to be independent from the subgoals that are 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   715
  produced by the rule. If we had applied it in the other order 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   716
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   717
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   718
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
   719
apply(tactic {* SUBGOAL select_tac 1 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   720
apply(tactic {* SUBGOAL select_tac 3 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   721
apply(tactic {* SUBGOAL select_tac 4 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   722
apply(tactic {* SUBGOAL select_tac 5 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   723
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   724
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   725
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   726
  then we have to be careful to not apply the tactic to the two subgoals produced by 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   727
  the first goal. To do this can result in quite messy code. In contrast, 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   728
  the ``reverse application'' is easy to implement.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   729
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   730
  Of course, this example is contrived: there are much simpler methods available in 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   731
  Isabelle for implementing a proof procedure analysing a goal according to its topmost
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   732
  connective. These simpler methods use tactic combinators, which we will explain 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   733
  in the next section.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   734
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   735
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   736
section {* Tactic Combinators *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   737
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   738
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   739
  The purpose of tactic combinators is to build compound tactics out of
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   740
  smaller tactics. In the previous section we already used @{ML THEN}, which
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   741
  just strings together two tactics in a sequence. For example:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   744
lemma shows "(Foo \<and> Bar) \<and> False"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   745
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
   746
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   747
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   748
       \end{minipage} *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   749
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   750
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   751
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   752
  If you want to avoid the hard-coded subgoal addressing, then you can use
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   753
  the ``primed'' version of @{ML THEN}. For example:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   754
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   756
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
   757
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
   758
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   759
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   760
       \end{minipage} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   763
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   764
  Here you only have to specify the subgoal of interest only once and
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   765
  it is consistently applied to the component tactics.
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   766
  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
   767
  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
   768
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   769
  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
   770
  sequence, you can use the combinator @{ML EVERY'}. For example
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   771
  the function @{ML foo_tac'} from page~\pageref{tac:footacprime} can also 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   772
  be written as:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   773
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   774
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   775
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
   776
                        atac, rtac @{thm disjI1}, atac]*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   777
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   778
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   779
  There is even another way of implementing this tactic: in automatic proof
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   780
  procedures (in contrast to tactics that might be called by the user) there
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   781
  are often long lists of tactics that are applied to the first
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   782
  subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' 1"}, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   783
  you can also just write
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   784
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   785
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   786
ML{*val foo_tac1 = EVERY1 [etac @{thm disjE}, rtac @{thm disjI2}, 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   787
                       atac, rtac @{thm disjI1}, atac]*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   788
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   789
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   790
  and call @{ML foo_tac1}.  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   791
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   792
  With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML EVERY1} it must be
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   793
  guaranteed that all component tactics successfully apply; otherwise the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   794
  whole tactic will fail. If you rather want to try out a number of tactics,
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   795
  then you can use the combinator @{ML ORELSE'} for two tactics, and @{ML
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   796
  FIRST'} (or @{ML FIRST1}) for a list of tactics. For example, the tactic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   797
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   798
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   799
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
   800
ML{*val orelse_xmp = rtac @{thm disjI1} ORELSE' rtac @{thm conjI}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   801
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   802
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   803
  will first try out whether rule @{text disjI} applies and after that 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   804
  @{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
   805
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   806
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   807
lemma shows "True \<and> False" and "Foo \<or> Bar"
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   808
apply(tactic {* orelse_xmp 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   809
apply(tactic {* orelse_xmp 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   810
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
   811
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   812
       \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   813
       @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   814
       \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   815
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   816
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   818
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   820
  Using @{ML FIRST'} we can simplify our @{ML select_tac} from Page~\pageref{tac:selecttac} 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   821
  as follows:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   822
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   823
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   824
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
   825
                          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
   826
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   827
text {*
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   828
  Since we like to mimic the behaviour of @{ML select_tac} as closely as possible, 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   829
  we must include @{ML all_tac} at the end of the list, otherwise the tactic will
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   830
  fail if no rule applies (we also have to wrap @{ML all_tac} using the 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   831
  @{ML K}-combinator, because it does not take a subgoal number as argument). You can 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   832
  test the tactic on the same goals:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   833
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   834
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   835
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
   836
apply(tactic {* select_tac' 4 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   837
apply(tactic {* select_tac' 3 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   838
apply(tactic {* select_tac' 2 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   839
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   840
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   841
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   842
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   843
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   844
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   845
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   846
  Since such repeated applications of a tactic to the reverse order of 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   847
  \emph{all} subgoals is quite common, there is the tactic combinator 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   848
  @{ML ALLGOALS} that simplifies this. Using this combinator you can simply 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   849
  write: *}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   850
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   851
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
   852
apply(tactic {* ALLGOALS select_tac' *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   853
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   854
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   855
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   856
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   857
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   858
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   859
  Remember that we chose to implement @{ML select_tac'} so that it 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   860
  always  succeeds. This can be potentially very confusing for the user, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   861
  for example,  in cases where the goal is the form
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   862
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   863
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   864
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   865
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   866
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   867
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   868
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   869
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   870
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   871
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   872
  In this case no rule applies. The problem for the user is that there is little 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   873
  chance to see whether or not progress in the proof has been made. By convention
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   874
  therefore, tactics visible to the user should either change something or fail.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   875
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   876
  To comply with this convention, we could simply delete the @{ML "K all_tac"}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   877
  from the end of the theorem list. As a result @{ML select_tac'} would only
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   878
  succeed on goals where it can make progress. But for the sake of argument,
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   879
  let us suppose that this deletion is \emph{not} an option. In such cases, you can
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   880
  use the combinator @{ML CHANGED} to make sure the subgoal has been changed
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   881
  by the tactic. Because now
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   882
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   883
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   884
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   885
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   886
apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   887
txt{* gives the error message:
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   888
  \begin{isabelle}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   889
  @{text "*** empty result sequence -- proof command failed"}\\
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   890
  @{text "*** At command \"apply\"."}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   891
  \end{isabelle} 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   892
*}(*<*)oops(*>*)
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   893
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   894
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   895
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   896
  We can further extend @{ML select_tac'} so that it not just applies to the topmost
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   897
  connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   898
  completely. For this you can use the tactic combinator @{ML REPEAT}. As an example 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   899
  suppose the following tactic
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   900
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   901
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   902
ML{*val repeat_xmp = REPEAT (CHANGED (select_tac' 1)) *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   903
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   904
text {* which applied to the proof *}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   905
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   906
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   907
apply(tactic {* repeat_xmp *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   908
txt{* produces
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   909
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   910
      \begin{minipage}{\textwidth}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   911
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   912
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   913
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   914
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   915
text {*
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   916
  Here it is crucial that @{ML select_tac'} is prefixed with @{ML CHANGED}, 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   917
  because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   918
  tactic as long as it succeeds). The function
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   919
  @{ML REPEAT1} is similar, but runs the tactic at least once (failing if 
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   920
  this is not possible).
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   921
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   922
  If you are after the ``primed'' version of @{ML repeat_xmp} then you 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   923
  need to implement it as
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   924
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   925
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   926
ML{*val repeat_xmp' = REPEAT o CHANGED o select_tac'*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   927
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   928
text {*
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   929
  since there are no ``primed'' versions of @{ML REPEAT} and @{ML CHANGED}.
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   930
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   931
  If you look closely at the goal state above, the tactics @{ML repeat_xmp}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   932
  and @{ML repeat_xmp'} are not yet quite what we are after: the problem is
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   933
  that goals 2 and 3 are not analysed. This is because the tactic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   934
  is applied repeatedly only to the first subgoal. To analyse also all
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   935
  resulting subgoals, you can use the tactic combinator @{ML REPEAT_ALL_NEW}. 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   936
  Suppose the tactic
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   937
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   938
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   939
ML{*val repeat_all_new_xmp = REPEAT_ALL_NEW (CHANGED o select_tac')*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   940
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   941
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   942
  you see that the following goal
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   943
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   944
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   945
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   946
apply(tactic {* repeat_all_new_xmp 1 *})
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   947
txt{* \begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   948
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   949
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   950
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   951
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   952
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   953
  is completely analysed according to the theorems we chose to
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   954
  include in @{ML select_tac'}. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   955
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   956
  Recall that tactics produce a lazy sequence of successor goal states. These
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   957
  states can be explored using the command \isacommand{back}. For example
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   958
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   959
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   960
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   961
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   962
apply(tactic {* etac @{thm disjE} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   963
txt{* applies the rule to the first assumption yielding the goal state:\smallskip
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   964
      
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   965
      \begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   966
      @{subgoals [display]}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   967
      \end{minipage}\smallskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   968
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   969
      After typing
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   970
  *}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   971
(*<*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   972
oops
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   973
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   974
apply(tactic {* etac @{thm disjE} 1 *})
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   975
(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   976
back
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   977
txt{* the rule now applies to the second assumption.\smallskip
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   978
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   979
      \begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   980
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   981
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   982
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   983
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   984
text {*
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   985
  Sometimes this leads to confusing behaviour of tactics and also has
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   986
  the potential to explode the search space for tactics.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   987
  These problems can be avoided by prefixing the tactic with the tactic 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   988
  combinator @{ML DETERM}.
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   989
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   990
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   991
lemma "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   992
apply(tactic {* DETERM (etac @{thm disjE} 1) *})
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   993
txt {*\begin{minipage}{\textwidth}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   994
      @{subgoals [display]}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   995
      \end{minipage} *}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   996
(*<*)oops(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   997
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   998
  This combinator will prune the search space to just the first successful application.
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   999
  Attempting to apply \isacommand{back} in this goal states gives the
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1000
  error message:
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1001
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1002
  \begin{isabelle}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1003
  @{text "*** back: no alternatives"}\\
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1004
  @{text "*** At command \"back\"."}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1005
  \end{isabelle}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1006
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1007
  \begin{readmore}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1008
  Most tactic combinators described in this section are defined in @{ML_file "Pure/tctical.ML"}.
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1009
  \end{readmore}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1010
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1011
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1012
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1013
section {* Rewriting and Simplifier Tactics *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1014
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1015
text {*
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1016
  @{ML rewrite_goals_tac}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1017
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
  @{ML ObjectLogic.full_atomize_tac}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1019
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
  @{ML ObjectLogic.rulify_tac}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1021
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1022
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1023
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1024
section {* Simprocs *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1025
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1026
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1027
  In Isabelle you can also implement custom simplification procedures, called
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1028
  \emph{simprocs}. Simprocs can be triggered on a specified term-pattern and
134
328370b75c33 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
  1029
  rewrite a term according to a theorem. They are useful in cases where
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1030
  a rewriting rule must be produced on ``demand'' or when rewriting by
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1031
  simplification is too unpredictable and potentially loops.
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1032
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1033
  To see how simprocs work, let us first write a simproc that just prints out
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1034
  the pattern which triggers it and otherwise does nothing. For this
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1035
  you can use the function:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1036
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1037
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1038
ML %linenosgray{*fun fail_sp_aux simpset redex = 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1039
let
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1040
  val ctxt = Simplifier.the_context simpset
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1041
  val _ = warning ("The redex: " ^ (str_of_cterm ctxt redex))
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1042
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1043
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1044
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1045
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1046
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1047
  This function takes a simpset and a redex (a @{ML_type cterm}) as
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1048
  arguments. In Lines 3 and~4, we first extract the context from the given
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1049
  simpset and then print out a message containing the redex.  The function
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1050
  returns @{ML NONE} (standing for an optional @{ML_type thm}) since at the
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1051
  moment we are \emph{not} interested in actually rewriting anything. We want
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1052
  that the simproc is triggered by the pattern @{term "Suc n"}. This can be
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1053
  done by adding the simproc to the current simproc as follows
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1054
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1055
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1056
simproc_setup fail_sp ("Suc n") = {* K fail_sp_aux *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1057
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1058
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1059
  where the second argument specifies the pattern and the right-hand side
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1060
  contains the code of the simproc (we have to use @{ML K} since we ignoring
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1061
  an argument about morphisms\footnote{FIXME: what does the morphism do?}). 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1062
  After this, the simplifier is aware of the simproc and you can test whether 
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1063
  it fires on the lemma:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1064
*}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1065
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1066
lemma shows "Suc 0 = 1"
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1067
  apply(simp)
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1068
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1069
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1070
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1071
  This will print out the message twice: once for the left-hand side and
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1072
  once for the right-hand side. The reason is that during simplification the
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1073
  simplifier will at some point reduce the term @{term "1::nat"} to @{term "Suc
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1074
  0"}, and then the simproc ``fires'' also on that term. 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1075
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1076
  We can add or delete the simproc from the current simpset by the usual 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1077
  \isacommand{declare}-statement. For example the simproc will be deleted
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1078
  with the declaration
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1079
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1080
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1081
declare [[simproc del: fail_sp]]
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1082
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1083
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1084
  If you want to see what happens with just \emph{this} simproc, without any 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1085
  interference from other rewrite rules, you can call @{text fail_sp} 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1086
  as follows:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1087
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1088
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1089
lemma shows "Suc 0 = 1"
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1090
  apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [@{simproc fail_sp}]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1091
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1092
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1093
text {*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1094
  Now the message shows up only once since the term @{term "1::nat"} is 
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1095
  left unchanged. 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1096
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1097
  Setting up a simproc using the command \isacommand{setup\_simproc} will
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1098
  always add automatically the simproc to the current simpset. If you do not
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1099
  want this, then you have to use a slightly different method for setting 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1100
  up the simproc. First the function @{ML fail_sp_aux} needs to be modified
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1101
  to
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1102
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1103
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1104
ML{*fun fail_sp_aux' simpset redex = 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1105
let
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1106
  val ctxt = Simplifier.the_context simpset
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1107
  val _ = warning ("The redex: " ^ (Syntax.string_of_term ctxt redex))
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1108
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1109
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1110
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1111
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1112
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1113
  Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1114
  (therefore we printing it out using the function @{ML string_of_term in Syntax}).
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1115
  We can turn this function into a proper simproc using
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1116
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1117
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1118
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1119
ML{*val fail_sp' = 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1120
  let 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1121
    val thy = @{theory}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1122
    val pat = [@{term "Suc n"}]
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1123
  in
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1124
    Simplifier.simproc_i thy "fail_sp'" pat (K fail_sp_aux')
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1125
  end*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1126
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1127
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1128
  Here the pattern is given as @{ML_type term} (instead of @{ML_type cterm}).
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1129
  The function also takes a list of patterns that can trigger the simproc.
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1130
  Now the simproc is set up and can be explicitly added using
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1131
  {ML addsimprocs} to a simpset whenerver
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1132
  needed. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1133
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1134
  Simprocs are applied from inside to outside and from left to right. You can
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1135
  see this in the proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1136
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1137
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1138
lemma shows "Suc (Suc 0) = (Suc 1)"
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1139
  apply(tactic {* simp_tac (HOL_ss addsimprocs [fail_sp']) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1140
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1141
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1142
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1143
  The simproc @{ML fail_sp'} prints out the sequence 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1144
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1145
@{text [display]
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1146
"> Suc 0
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1147
> Suc (Suc 0) 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1148
> Suc 1"}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1149
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1150
  To see how a simproc applies a theorem,  let us implement a simproc that
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1151
  rewrites terms according to the equation:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1152
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1153
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1154
lemma plus_one: 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1155
  shows "Suc n \<equiv> n + 1" by simp
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1156
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1157
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1158
  Simprocs expect that the given equation is a meta-equation, however the
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1159
  equation can contain preconditions (the simproc then will only fire if the
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1160
  preconditions can be solved). To see that one has relatively precise control over
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1161
  the rewriting with simprocs, let us further assume we want that the simproc
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1162
  only rewrites terms ``greater'' than @{term "Suc 0"}. For this we can write 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1163
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1164
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1165
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1166
ML{*fun plus_one_sp_aux ss redex =
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1167
  case redex of
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1168
    @{term "Suc 0"} => NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1169
  | _ => SOME @{thm plus_one}*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1170
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1171
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1172
  and set up the simproc as follows.
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1173
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1174
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1175
ML{*val plus_one_sp =
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1176
  let
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1177
    val thy = @{theory}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1178
    val pat = [@{term "Suc n"}] 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1179
  in
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1180
    Simplifier.simproc_i thy "sproc +1" pat (K plus_one_sp_aux)
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1181
  end*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1182
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1183
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1184
  Now the simproc is set up so that it is triggered by terms
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1185
  of the form @{term "Suc n"}, but inside the simproc we only produce
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1186
  a theorem if the term is not @{term "Suc 0"}. The result you can see
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1187
  in the following proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1188
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1189
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1190
lemma shows "P (Suc (Suc (Suc 0))) (Suc 0)"
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1191
  apply(tactic {* simp_tac (HOL_ss addsimprocs [plus_one_sp]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1192
txt{*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1193
  where the simproc produces the goal state
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1194
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1195
  @{subgoals[display]}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1196
*}  
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1197
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1198
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1199
text {*
133
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
  1200
  As usual with rewriting you have to worry about looping: you already have
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1201
  a loop with @{ML plus_one_sp}, if you apply it with the default simpset (because
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1202
  the default simpset contains a rule which just does the opposite of @{ML plus_one_sp},
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1203
  namely rewriting @{text [quotes] "+ 1"} to a successor). So you have to be careful 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1204
  in choosing the right simpset to which you add a simproc. 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1205
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1206
  Next let us implement a simproc that replaces terms of the form @{term "Suc n"}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1207
  with the number @{text n} increase by one. First we implement a function that
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1208
  takes a term and produces the corresponding integer value.
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1209
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1210
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1211
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1212
ML{*fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1213
  | dest_suc_trm t = snd (HOLogic.dest_number t)*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1214
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1215
text {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1216
  It uses the library function @{ML dest_number in HOLogic} that transforms
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1217
  (Isabelle) terms, like @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1218
  on, into integer values. This function raises the exception @{ML TERM}, if
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1219
  the term is not a number. The next function expects a pair consisting of a term
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1220
  @{text t} (containing @{term Suc}s) and the corresponding integer value @{text n}. 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1221
*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1222
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1223
ML %linenosgray{*fun get_thm ctxt (t, n) =
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1224
let
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1225
  val num = HOLogic.mk_number @{typ "nat"} n
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1226
  val goal = Logic.mk_equals (t, num)
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1227
in
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1228
  Goal.prove ctxt [] [] goal (K (arith_tac ctxt 1))
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1229
end*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1230
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1231
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1232
  From the integer value it generates the corresponding number term, called 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1233
  @{text num} (Line 3), and then generates the meta-equation @{text "t \<equiv> num"} 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1234
  (Line 4), which it proves by the arithmetic tactic in Line 6. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1235
134
328370b75c33 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
  1236
  For our purpose at the moment, proving the meta-equation using @{ML arith_tac} is
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1237
  fine, but there is also an alternative employing the simplifier with a very
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1238
  restricted simpset. For the kind of lemmas we want to prove, the simpset
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1239
  @{text "num_ss"} in the code will suffice.
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1240
*}
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1241
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1242
ML{*fun get_thm_alt ctxt (t, n) =
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1243
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1244
  val num = HOLogic.mk_number @{typ "nat"} n
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1245
  val goal = Logic.mk_equals (t, num)
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1246
  val num_ss = HOL_ss addsimps [@{thm One_nat_def}, @{thm Let_def}] @ 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1247
          @{thms nat_number} @ @{thms neg_simps} @ @{thms plus_nat.simps}
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1248
in
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1249
  Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1))
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1250
end*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1251
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1252
text {*
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1253
  The advantage of @{ML get_thm_alt} is that it leaves very little room for 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1254
  something to go wrong; in contrast it is much more difficult to predict 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1255
  what happens with @{ML arith_tac}, especially in more complicated 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1256
  circumstances. The disatvantage of @{ML get_thm_alt} is to find a simpset
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1257
  that is sufficiently powerful to solve every instance of the lemmas
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1258
  we like to prove. This requires careful tuning, but is often necessary in 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1259
  ``production code''.\footnote{It would be of great help if there is another
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1260
  way than tracing the simplifier to obtain the lemmas that are successfully 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1261
  applied during simplification. Alas, there is none.} 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1262
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1263
  Anyway, either version can be used in the function that produces the actual 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1264
  theorem for the simproc.
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1265
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1266
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1267
ML{*fun nat_number_sp_aux ss t =
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1268
let 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1269
  val ctxt = Simplifier.the_context ss
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1270
in
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1271
  SOME (get_thm ctxt (t, dest_suc_trm t))
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1272
  handle TERM _ => NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1273
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1274
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1275
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1276
  This function uses the fact that @{ML dest_suc_trm} might throw an exception 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1277
  @{ML TERM}. In this case there is nothing that can be rewritten and therefore no
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1278
  theorem is produced (i.e.~the function returns @{ML NONE}). To try out the simproc 
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1279
  on an example, you can set it up as follows:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1280
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1281
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1282
ML{*val nat_number_sp =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1283
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1284
  val thy = @{theory}
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1285
  val pat = [@{term "Suc n"}]
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1286
in 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1287
  Simplifier.simproc_i thy "nat_number" pat (K nat_number_sp_aux) 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1288
end*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1289
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1290
text {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1291
  Now in the lemma
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1292
  *}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1293
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1294
lemma "P (Suc (Suc 2)) (Suc 99) (0::nat) (Suc 4 + Suc 0) (Suc (0 + 0))"
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1295
  apply(tactic {* simp_tac (HOL_ss addsimprocs [nat_number_sp]) 1*})
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1296
txt {* 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1297
  you obtain the more legible goal state
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1298
  
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1299
  @{subgoals [display]}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1300
  *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1301
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1302
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1303
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1304
  where the simproc rewrites all @{term "Suc"}s except in the last argument. There it cannot 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1305
  rewrite anything, because it does not know how to transform the term @{term "Suc (0 + 0)"}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1306
  into a number. To solve this problem have a look at the next exercise. 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1307
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1308
  \begin{exercise}\label{ex:addsimproc}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1309
  Write a simproc that replaces terms of the form @{term "t\<^isub>1 + t\<^isub>2"} by their 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1310
  result. You can assume the terms are ``proper'' numbers, that is of the form
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1311
  @{term "0::nat"}, @{term "1::nat"}, @{term "2::nat"} and so on.
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1312
  \end{exercise}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1313
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1314
  (FIXME: We did not do anything with morphisms. Anything interesting
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1315
  one can say about them?)
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1316
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1317
137
a9685909944d new pfd file
Christian Urban <urbanc@in.tum.de>
parents: 135
diff changeset
  1318
section {* Conversions\label{sec:conversion} *}
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1319
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1320
text {*
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1321
Conversions are meta-equalities depending on some input term. Their type is
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1322
as follows:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1323
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1324
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1325
ML {*type conv = Thm.cterm -> Thm.thm*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1326
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1327
text {*
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1328
The simplest two conversions are @{ML "Conv.all_conv"}, which maps a term to an instance of the reflexivity theorem, and @{ML "Conv.no_conv"}, which always fails:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1329
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1330
@{ML_response_fake "Conv.all_conv @{cterm True}" "True \<equiv> True"}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1331
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1332
@{ML_response_fake "Conv.no_conv @{cterm True}" "*** Exception- CTERM (\"no conversion\", []) raised"}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1333
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1334
A further basic conversion is, for example, @{ML "Thm.beta_conversion"}:
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1335
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1336
@{ML_response_fake "Thm.beta_conversion true @{cterm \"(\<lambda>x. x \<or> False) True\"}"
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1337
"(\<lambda>x. x \<or> False) True \<equiv> True \<or> False"}
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1338
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1339
User-defined rewrite rules can be applied by the conversion
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1340
@{ML "Conv.rewr_conv"}. Consider, for example, the following rule:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1341
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1342
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1343
lemma true_conj1: "True \<and> P \<equiv> P" by simp
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1344
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1345
text {*
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1346
Here is how this rule can be used for rewriting:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1347
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1348
@{ML_response_fake "Conv.rewr_conv @{thm true_conj1} @{cterm \"True \<or> False\"}"
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1349
 "True \<and> False \<equiv> False"}
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1350
*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1351
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1352
text {*
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1353
Basic conversions can be combined with a number of conversionals, i.e.
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1354
conversion combinators:
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1355
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1356
@{ML_response_fake
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1357
"(Thm.beta_conversion true then_conv Conv.rewr_conv @{thm true_conj1})
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1358
  @{cterm \"(\<lambda>x. x \<and> False) True\"}"
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1359
"(\<lambda>x. x \<and> False) True \<equiv> False"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1360
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1361
@{ML_response_fake
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1362
"(Conv.rewr_conv @{thm true_conj1} else_conv Conv.all_conv)
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1363
  @{cterm \"P \<or> (True \<and> Q)\"}"
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1364
"P \<or> (True \<and> Q) \<equiv> P \<or> (True \<and> Q)"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1365
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1366
@{ML_response_fake
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1367
"Conv.arg_conv (Conv.rewr_conv @{thm true_conj1})
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1368
  @{cterm \"P \<or> (True \<and> Q)\"}"
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1369
"P \<or> (True \<and> Q) \<equiv> P \<or> Q"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1370
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1371
\begin{readmore}
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1372
See @{ML_file "Pure/conv.ML"} for more conversionals. Further basic conversions
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1373
can be found in, for example, @{ML_file "Pure/thm.ML"},
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1374
@{ML_file "Pure/drule.ML"}, and @{ML_file "Pure/meta_simplifier.ML"}.
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1375
\end{readmore}
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1376
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1377
Conversions are a thin layer on top of Isabelle's inference kernel, and may
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1378
be seen as a controllable, bare-bone version of Isabelle's simplifier. We
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1379
will demonstrate this feature in the following example.
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1380
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1381
To begin with, let's assumes we want to simplify with the rewrite rule
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1382
@{text true_conj1}. As a conversion, this may look as follows:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1383
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1384
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1385
ML {*fun all_true1_conv ctxt ct =
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1386
  (case Thm.term_of ct of
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1387
    @{term "op \<and>"} $ @{term True} $ _ => 
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1388
      (Conv.arg_conv (all_true1_conv ctxt) then_conv
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1389
        Conv.rewr_conv @{thm true_conj1}) ct
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1390
  | _ $ _ => Conv.combination_conv (all_true1_conv ctxt)
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1391
      (all_true1_conv ctxt) ct
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1392
  | Abs _ => Conv.abs_conv (fn (_, cx) => all_true1_conv cx) ctxt ct
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1393
  | _ => Conv.all_conv ct)*}
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1394
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1395
text {*
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1396
Here is this conversion in action:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1397
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1398
@{ML_response_fake
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1399
"all_true1_conv @{context}
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1400
  @{cterm \"distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x\"}"
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1401
"distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x \<equiv> distinct [1, x] \<longrightarrow> 1 \<noteq> x"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1402
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1403
Now, let's complicate the task a bit: Rewrite according to the rule
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1404
@{text true_conj1}, but only in the first arguments of @{term If}:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1405
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1406
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1407
ML {*fun if_true1_conv ctxt ct =
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1408
  (case Thm.term_of ct of
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1409
    Const (@{const_name If}, _) $ _ =>
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1410
      Conv.arg_conv (all_true1_conv ctxt) ct
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1411
  | _ $ _ => Conv.combination_conv (if_true1_conv ctxt)
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1412
      (if_true1_conv ctxt) ct
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1413
  | Abs _ => Conv.abs_conv (fn (_, cx) => if_true1_conv cx) ctxt ct
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1414
  | _ => Conv.all_conv ct)*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1415
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1416
text {*
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1417
Here is an application of this conversion:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1418
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1419
@{ML_response_fake
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1420
"if_true1_conv @{context}
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1421
  @{cterm \"if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False\"}"
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1422
"if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False \<equiv> if P (1 \<noteq> 2) then True \<and> True else True \<and> False"}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1423
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1424
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1425
text {*
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1426
Conversions can also be turned into a tactic with the @{ML CONVERSION}
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1427
tactical, and there are predefined conversionals to traverse a goal state.
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1428
Given a state @{term "\<And>x. P \<Longrightarrow> Q"},  the conversional
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1429
@{ML Conv.params_conv} applies a conversion to @{term "P \<Longrightarrow> Q"};
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1430
given a state @{term "\<lbrakk> P1; P2 \<rbrakk> \<Longrightarrow> Q"},
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1431
the conversional @{ML Conv.prems_conv} applies a conversion to the premises
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1432
@{term P1} and @{term P2}, and @{ML Conv.concl_conv} applies a conversion to
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1433
the conclusion @{term Q}.
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1434
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1435
Assume we want to apply @{ML all_true1_conv} only in the conclusion
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1436
of the goal, and @{ML if_true1_conv} should only be applied in the premises.
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1437
Here is a tactic doing exactly that:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1438
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1439
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1440
ML {*val true1_tac = CSUBGOAL (fn (goal, i) =>
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1441
  let val ctxt = ProofContext.init (Thm.theory_of_cterm goal)
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1442
  in
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1443
    CONVERSION (
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1444
      Conv.params_conv ~1 (fn cx =>
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1445
        (Conv.prems_conv ~1 (if_true1_conv cx)) then_conv
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1446
          Conv.concl_conv ~1 (all_true1_conv cx)) ctxt) i
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1447
  end)*}
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1448
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1449
text {* 
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1450
To demonstrate this tactic, consider the following example:
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1451
*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1452
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1453
lemma
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1454
  "if True \<and> P then P else True \<and> False \<Longrightarrow>
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1455
  (if True \<and> Q then Q else P) \<longrightarrow> True \<and> (True \<and> Q)"
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1456
apply(tactic {* true1_tac 1 *})
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1457
txt {* \begin{minipage}{\textwidth}
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1458
       @{subgoals [display]} 
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1459
       \end{minipage} *}
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1460
(*<*)oops(*>*)
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1461
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1462
text {*
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1463
Conversions are not restricted to work on certified terms only, they can also
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1464
be lifted to theorem transformers, i.e. functions mapping a theorem to a
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1465
theorem, by the help of @{ML Conv.fconv_rule}. As an example, consider the
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1466
conversion @{ML all_true1_conv} again:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1467
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1468
@{ML_response_fake
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1469
"Conv.fconv_rule (all_true1_conv @{context})
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1470
  @{lemma \"P \<or> (True \<and> \<not>P)\" by simp}" "P \<or> \<not>P"}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1471
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1472
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1473
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1474
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1475
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1476
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1477
section {* Structured Proofs *}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1478
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1479
text {* TBD *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1480
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1481
lemma True
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1482
proof
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1483
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1484
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1485
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1486
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1487
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1488
    then have "A & B" ..
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1489
    then have C by (rule r)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1490
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1491
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1492
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1493
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1494
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1495
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1496
    note conjI [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1497
    note r [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1498
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1499
oops
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1500
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1501
ML {* fun prop ctxt s =
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1502
  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
  1503
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1504
ML {* 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1505
  val ctxt0 = @{context};
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1506
  val ctxt = ctxt0;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1507
  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
  1508
  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
  1509
  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
  1510
  val this = [@{thm conjI} OF this]; 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1511
  val this = r OF this;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1512
  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
  1513
  val this = Variable.export ctxt ctxt0 [this] 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  1514
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1515
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1516
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
  1517
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1518
end