CookBook/Tactical.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 05 Mar 2009 16:46:43 +0000
changeset 160 cc9359bfacf4
parent 158 d7944bdf7b3f
child 161 83f36a1c62f2
permissions -rw-r--r--
redefined the functions warning and tracing in order to properly match more antiquotations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Tactical
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
     2
imports Base FirstSteps
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
chapter {* Tactical Reasoning\label{chp:tactical} *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
     8
  The main reason for descending to the ML-level of Isabelle is to be able to
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
     9
  implement automatic proof procedures. Such proof procedures usually lessen
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    10
  considerably the burden of manual reasoning, for example, when introducing
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    11
  new definitions. These proof procedures are centred around refining a goal
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    12
  state using tactics. This is similar to the \isacommand{apply}-style
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
    13
  reasoning at the user-level, where goals are modified in a sequence of proof
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    14
  steps until all of them are solved. However, there are also more structured
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    15
  operations available on the ML-level that help with the handling of
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    16
  variables and assumptions.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    17
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    20
section {* Basics of Reasoning with Tactics*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
text {*
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    23
  To see how tactics work, let us first transcribe a simple \isacommand{apply}-style proof 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
    24
  into ML. Suppose the following proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
apply(erule disjE)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
apply(rule disjI2)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
apply(rule disjI1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  This proof translates to the following ML-code.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
@{ML_response_fake [display,gray]
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
"let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  Goal.prove ctxt [\"P\", \"Q\"] [] goal 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
   (fn _ => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
      etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
      THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
      THEN atac 1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  To start the proof, the function @{ML "Goal.prove"}~@{text "ctxt xs As C
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    53
  tac"} sets up a goal state for proving the goal @{text C} 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    54
  (that is @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    55
  assumptions @{text As} (happens to be empty) with the variables
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  @{text xs} that will be generalised once the goal is proved (in our case
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  @{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  it can make use of the local assumptions (there are none in this example). 
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    59
  The functions @{ML etac}, @{ML rtac} and @{ML atac} in the code above correspond to 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  @{text erule}, @{text rule} and @{text assumption}, respectively. 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    61
  The operator @{ML THEN} strings the tactics together. 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  \begin{readmore}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    64
  To learn more about the function @{ML Goal.prove} see \isccite{sec:results}
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
    65
  and the file @{ML_file "Pure/goal.ML"}.  See @{ML_file
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    66
  "Pure/tactic.ML"} and @{ML_file "Pure/tctical.ML"} for the code of basic
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    67
  tactics and tactic combinators; see also Chapters 3 and 4 in the old
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    68
  Isabelle Reference Manual.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
    71
  Note that in the code above we use antiquotations for referencing the theorems. Many theorems
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    72
  also have ML-bindings with the same name. Therefore, we could also just have
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
    73
  written @{ML "etac disjE 1"}, or in case where there are no ML-binding obtain
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    74
  the theorem dynamically using the function @{ML thm}; for example 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
    75
  \mbox{@{ML  "etac (thm \"disjE\") 1"}}. Both ways however are considered bad style! 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    76
  The reason
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    77
  is that the binding for @{ML disjE} can be re-assigned by the user and thus
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    78
  one does not have complete control over which theorem is actually
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    79
  applied. This problem is nicely prevented by using antiquotations, because
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
    80
  then the theorems are fixed statically at compile-time.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    82
  During the development of automatic proof procedures, you will often find it 
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
    83
  necessary to test a tactic on examples. This can be conveniently 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  Consider the following sequence of tactics
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
ML{*val foo_tac = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
      (etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
       THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
       THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
       THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
       THEN atac 1)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
text {* and the Isabelle proof: *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
lemma "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
apply(tactic {* foo_tac *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   102
  By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} you can call from the 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
   103
  user-level of Isabelle the tactic @{ML foo_tac} or 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   104
  any other function that returns a tactic.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   106
  The tactic @{ML foo_tac} is just a sequence of simple tactics stringed
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   107
  together by @{ML THEN}. As can be seen, each simple tactic in @{ML foo_tac}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   108
  has a hard-coded number that stands for the subgoal analysed by the
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   109
  tactic (@{text "1"} stands for the first, or top-most, subgoal). This hard-coding
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   110
  of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   111
  you can write\label{tac:footacprime}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
ML{*val foo_tac' = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
      (etac @{thm disjE} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
       THEN' rtac @{thm disjI2} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
       THEN' atac 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
       THEN' rtac @{thm disjI1} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
       THEN' atac)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
text {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  and then give the number for the subgoal explicitly when the tactic is
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   123
  called. So in the next proof you can first discharge the second subgoal, and
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   124
  subsequently the first.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
   and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
apply(tactic {* foo_tac' 2 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
text {*
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   134
  This kind of addressing is more difficult to achieve when the goal is 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   135
  hard-coded inside the tactic. For most operators that combine tactics 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   136
  (@{ML THEN} is only one such operator) a ``primed'' version exists.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   137
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   138
  The tactics @{ML foo_tac} and @{ML foo_tac'} are very specific for
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   139
  analysing goals being only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   140
  of this form, then they return the error message:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   141
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   142
  \begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   143
  @{text "*** empty result sequence -- proof command failed"}\\
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   144
  @{text "*** At command \"apply\"."}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   145
  \end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   146
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   147
  This means the tactics failed. The reason for this error message is that tactics 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   148
  are functions mapping a goal state to a (lazy) sequence of successor states. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   149
  Hence the type of a tactic is:
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   150
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   152
ML{*type tactic = thm -> thm Seq.seq*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   154
text {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   155
  By convention, if a tactic fails, then it should return the empty sequence. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   156
  Therefore, if you write your own tactics, they  should not raise exceptions 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   157
  willy-nilly; only in very grave failure situations should a tactic raise the 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   158
  exception @{text THM}.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   159
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   160
  The simplest tactics are @{ML no_tac} and @{ML all_tac}. The first returns
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   161
  the empty sequence and is defined as
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   162
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   163
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   164
ML{*fun no_tac thm = Seq.empty*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   165
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   166
text {*
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   167
  which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   168
  up in a single member sequence; it is defined as
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   169
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   170
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   171
ML{*fun all_tac thm = Seq.single thm*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   172
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   173
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   174
  which means @{ML all_tac} always succeeds, but also does not make any progress 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   175
  with the proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   177
  The lazy list of possible successor goal states shows through at the user-level 
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   178
  of Isabelle when using the command \isacommand{back}. For instance in the 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   179
  following proof there are two possibilities for how to apply 
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   180
  @{ML foo_tac'}: either using the first assumption or the second.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
back
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   188
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   190
  By using \isacommand{back}, we construct the proof that uses the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   191
  second assumption. While in the proof above, it does not really matter which 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   192
  assumption is used, in more interesting cases provability might depend
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   193
  on exploring different possibilities.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   194
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   197
  sequences. In day-to-day Isabelle programming, however, one rarely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   198
  constructs sequences explicitly, but uses the predefined tactics and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   199
  tactic combinators instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   202
  It might be surprising that tactics, which transform
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   203
  one goal state to the next, are functions from theorems to theorem 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   204
  (sequences). The surprise resolves by knowing that every 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   205
  goal state is indeed a theorem. To shed more light on this,
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   206
  let us modify the code of @{ML all_tac} to obtain the following
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   207
  tactic
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   208
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   209
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   210
ML{*fun my_print_tac ctxt thm =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   211
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   212
  val _ = warning (str_of_thm ctxt thm)
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   213
in 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   214
  Seq.single thm
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   215
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
   216
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   217
text_raw {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   218
  \begin{figure}[p]
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   219
  \begin{isabelle}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   220
*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   221
lemma shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B" 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   222
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   223
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   224
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   225
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   226
      \end{minipage}\medskip      
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   227
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   228
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   229
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   230
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   231
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   232
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   233
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   234
      \end{minipage}\medskip
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   237
apply(rule conjI)
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   238
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   239
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   240
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   241
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   242
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   243
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   244
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   245
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   246
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   247
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   248
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   249
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   250
      \end{minipage}\medskip
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   251
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   252
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   253
apply(assumption)
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   254
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   255
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   256
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   257
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   258
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   259
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   260
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   261
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   262
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   263
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   264
      @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   265
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   266
      \end{minipage}\medskip
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   267
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   268
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   269
apply(assumption)
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   270
apply(tactic {* my_print_tac @{context} *})
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   271
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   272
txt{* \begin{minipage}{\textwidth}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   273
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   274
      \end{minipage}\medskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   275
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   276
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   277
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   278
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   279
      internal goal state:\\
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   280
      @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   281
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   282
      \end{minipage}\medskip   
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   283
   *}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   284
done
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   285
text_raw {*  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   286
  \end{isabelle}
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   287
  \caption{The figure shows a proof where each intermediate goal state is
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   288
  printed by the Isabelle system and by @{ML my_print_tac}. The latter shows
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   289
  the goal state as represented internally (highlighted boxes). This
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   290
  illustrates that every goal state in Isabelle is represented by a theorem:
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   291
  when you start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}} the theorem is
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   292
  @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> (\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when you finish the proof the
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   293
  theorem is @{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}.\label{fig:goalstates}}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   294
  \end{figure}
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   295
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   296
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   297
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   298
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   299
  which prints out the given theorem (using the string-function defined in
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   300
  Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   301
  tactic we are in the position to inspect every goal state in a
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   302
  proof. Consider now the proof in Figure~\ref{fig:goalstates}: as can be seen, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   303
  internally every goal state is an implication of the form
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   304
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   305
  @{text[display] "A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> (C)"}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   306
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   307
  where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   308
  the subgoals. So after setting up the lemma, the goal state is always of the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   309
  form @{text "C \<Longrightarrow> (C)"}; when the proof is finished we are left with @{text
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   310
  "(C)"}. Since the goal @{term C} can potentially be an implication, there is
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   311
  a ``protector'' wrapped around it (in from of an outermost constant @{text
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   312
  "Const (\"prop\", bool \<Rightarrow> bool)"}; however this constant
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   313
  is invisible in the figure). This wrapper prevents that premises of @{text C} are
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   314
  mis-interpreted as open subgoals. While tactics can operate on the subgoals
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   315
  (the @{text "A\<^isub>i"} above), they are expected to leave the conclusion
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   316
  @{term C} intact, with the exception of possibly instantiating schematic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   317
  variables. If you use the predefined tactics, which we describe in the next
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   318
  section, this will always be the case.
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   319
 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   320
  \begin{readmore}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   321
  For more information about the internals of goals see \isccite{sec:tactical-goals}.
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   322
  \end{readmore}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   323
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   324
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   325
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   326
section {* Simple Tactics *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   328
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   329
  Let us start with the tactic @{ML print_tac}, which is quite useful for low-level 
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   330
  debugging of tactics. It just prints out a message and the current goal state 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
   331
  (unlike @{ML my_print_tac}, it prints the goal state as the user would see it). 
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   332
  For example, 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 
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   389
  basic tactics take such a number as argument; it addresses the subgoal they are analysing. 
109
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 {*
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   456
  where the application of rule @{text bspec} generates two subgoals involving the
109
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
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   481
  To improve readability of the theorems we produce below, we shall use the
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   482
  function @{ML no_vars} from Section~\ref{sec:printing}, which transforms
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   483
  schematic variables into free ones.  Using this function for the first @{ML
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   484
  RS}-expression above produces the more readable result:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   485
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   486
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   487
  "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
   488
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   489
  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
   490
  the function @{ML MRS}:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   491
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   492
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   493
  "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
   494
  "\<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
   495
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   496
  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
   497
  functions @{ML RL} and @{ML MRL}. For example in the code below,
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   498
  every theorem in the second list is instantiated with every 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   499
  theorem in the first.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   500
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   501
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   502
   "[@{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
   503
"[\<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
   504
 \<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
   505
 (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
   506
 Q \<Longrightarrow> (P \<or> Q) \<or> Qa]"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   507
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   508
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   509
  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
   510
  \end{readmore}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   511
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   512
  Often proofs on the ML-level involve elaborate operations on assumptions and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   513
  @{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
   514
  shown so far is very unwieldy and brittle. Some convenience and
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   515
  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
   516
  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
   517
  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
   518
  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
   519
  string transformation functions from in Section~\ref{sec:printing}). Consider
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   520
  now the proof:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   521
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   522
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   523
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   524
\begin{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   525
\begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   526
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   527
ML{*fun sp_tac {prems, params, asms, concl, context, schematics} = 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   528
let 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   529
  val str_of_params = str_of_cterms context params
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   530
  val str_of_asms = str_of_cterms context asms
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   531
  val str_of_concl = str_of_cterm context concl
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   532
  val str_of_prems = str_of_thms context prems   
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   533
  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
   534
 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   535
  val _ = (warning ("params: " ^ str_of_params);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   536
           warning ("schematics: " ^ str_of_schms);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   537
           warning ("assumptions: " ^ str_of_asms);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   538
           warning ("conclusion: " ^ str_of_concl);
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   539
           warning ("premises: " ^ str_of_prems))
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   540
in
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   541
  no_tac 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   542
end*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   543
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   544
\end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   545
\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
   546
 @{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
   547
  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
   548
\end{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   549
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   550
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   551
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   552
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
   553
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   554
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   555
txt {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   556
  The tactic produces the following printout:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   557
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   558
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   559
  \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
   560
  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
   561
  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
   562
  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
   563
  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
   564
  premises:    & @{term "A x y"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   565
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   566
  \end{quote}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   567
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   568
  Notice 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
   569
  @{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
   570
  the subproof. By convention these fixed variables are printed in brown colour.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   571
  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
   572
  @{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
   573
  @{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
   574
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   575
  Notice also that we had to append @{text [quotes] "?"} to the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   576
  \isacommand{apply}-command. The reason is that @{ML SUBPROOF} normally
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   577
  expects that the subgoal is solved completely.  Since in the function @{ML
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   578
  sp_tac} we returned the tactic @{ML no_tac}, the subproof obviously
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   579
  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
   580
  graceful manner so that the warning messages are not overwritten by an 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   581
  ``empty sequence'' error message.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   582
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   583
  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
   584
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   585
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   586
apply(rule impI)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   587
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   588
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   589
txt {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   590
  then the tactic prints out: 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   591
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   592
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   593
  \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
   594
  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
   595
  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
   596
  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
   597
  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
   598
  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
   599
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   600
  \end{quote}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   601
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   602
(*<*)oops(*>*)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   603
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   604
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   605
  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
   606
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   607
  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
   608
  using the usual tactics, because the parameter @{text prems} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   609
  contains them as theorems. With this you can easily 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   610
  implement a tactic that behaves almost like @{ML atac}:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   611
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   612
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   613
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
   614
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   615
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   616
  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
   617
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   618
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   619
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
   620
apply(tactic {* atac' @{context} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   621
txt{* it will produce
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   622
      @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   623
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   624
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   625
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   626
  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
   627
  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
   628
  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
   629
  an advantage in the situations for which @{ML SUBPROOF} was designed: 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   630
  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
   631
  several goals and can render them unprovable. @{ML SUBPROOF} is meant 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   632
  to avoid this.
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   633
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   634
  Notice that @{ML atac'} inside @{ML SUBPROOF} calls @{ML resolve_tac} with 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   635
  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
   636
  the \isacommand{apply}-step uses @{text "1"}. This is another advantage 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   637
  of @{ML SUBPROOF}: the addressing  inside it is completely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   638
  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
   639
  @{ML atac'} to the second goal by just writing:
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   640
*}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   641
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   642
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
   643
apply(tactic {* atac' @{context} 2 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   644
apply(rule TrueI)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   645
done
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   646
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   647
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   649
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   650
  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
   651
  also described in \isccite{sec:results}. 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   652
  \end{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   653
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   654
  Similar but less powerful functions than @{ML SUBPROOF} are @{ML SUBGOAL}
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   655
  and @{ML CSUBGOAL}. They allow you to inspect a given subgoal (the former
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   656
  presents the subgoal as a @{ML_type term}, while the latter as a @{ML_type
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   657
  cterm}). With this you can implement a tactic that applies a rule according
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   658
  to the topmost logic connective in the subgoal (to illustrate this we only
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   659
  analyse a few connectives). The code of the tactic is as
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   660
  follows.\label{tac:selecttac}
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   661
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   664
ML %linenosgray{*fun select_tac (t, i) =
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   665
  case t of
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   666
     @{term "Trueprop"} $ t' => select_tac (t', i)
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   667
   | @{term "op \<Longrightarrow>"} $ _ $ t' => select_tac (t', i)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   668
   | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   669
   | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   670
   | @{term "Not"} $ _ => rtac @{thm notI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   671
   | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   672
   | _ => all_tac*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   673
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   674
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   675
  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
   676
  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
   677
  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
   678
  analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   679
  analysed. Similarly with meta-implications in the next line.  While for the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   680
  first five patterns we can use the @{text "@term"}-antiquotation to
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   681
  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
   682
  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
   683
  quantified variable. So you really have to construct the pattern using the
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   684
  basic term-constructors. This is not necessary in other cases, because their
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   685
  type is always fixed to function types involving only the type @{typ
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   686
  bool}. (See Section \ref{sec:terms_types_manually} about constructing terms
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   687
  manually.) For the catch-all pattern, we chose to just return @{ML all_tac}. 
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   688
  Consequently, @{ML select_tac} never fails.
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   689
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   690
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   691
  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
   692
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   693
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   694
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   695
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
   696
apply(tactic {* SUBGOAL select_tac 4 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   697
apply(tactic {* SUBGOAL select_tac 3 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   698
apply(tactic {* SUBGOAL select_tac 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   699
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
   700
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   701
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   702
      \end{minipage} *}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   703
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   704
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   705
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   706
  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
   707
  Note that we applied the tactic to the goals in ``reverse'' order. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   708
  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
   709
  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
   710
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   711
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   712
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
   713
apply(tactic {* SUBGOAL select_tac 1 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   714
apply(tactic {* SUBGOAL select_tac 3 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   715
apply(tactic {* SUBGOAL select_tac 4 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   716
apply(tactic {* SUBGOAL select_tac 5 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   717
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   718
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   719
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   720
  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
   721
  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
   722
  the ``reverse application'' is easy to implement.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   723
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   724
  Of course, this example is
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   725
  contrived: there are much simpler methods available in Isabelle for
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   726
  implementing a proof procedure analysing a goal according to its topmost
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   727
  connective. These simpler methods use tactic combinators, which we will
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   728
  explain in the next section.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
   729
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   730
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   731
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   732
section {* Tactic Combinators *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   733
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   734
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   735
  The purpose of tactic combinators is to build compound tactics out of
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   736
  smaller tactics. In the previous section we already used @{ML THEN}, which
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   737
  just strings together two tactics in a sequence. For example:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   740
lemma shows "(Foo \<and> Bar) \<and> False"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   741
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
   742
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   743
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   744
       \end{minipage} *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   745
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   746
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   747
text {*
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   748
  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
   749
  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
   750
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   752
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
   753
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
   754
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   755
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   756
       \end{minipage} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   759
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   760
  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
   761
  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
   762
  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
   763
  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
   764
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   765
  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
   766
  sequence, you can use the combinator @{ML EVERY'}. For example
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   767
  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
   768
  be written as:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   769
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   770
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   771
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
   772
                        atac, rtac @{thm disjI1}, atac]*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   773
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   774
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   775
  There is even another way of implementing this tactic: in automatic proof
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   776
  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
   777
  are often long lists of tactics that are applied to the first
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   778
  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
   779
  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
   780
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   781
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   782
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
   783
                       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
   784
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   785
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   786
  and call @{ML foo_tac1}.  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   787
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   788
  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
   789
  guaranteed that all component tactics successfully apply; otherwise the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   790
  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
   791
  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
   792
  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
   793
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   794
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   795
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
   796
ML{*val orelse_xmp = rtac @{thm disjI1} ORELSE' rtac @{thm conjI}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   797
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   798
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   799
  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
   800
  @{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
   801
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   802
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   803
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
   804
apply(tactic {* orelse_xmp 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   805
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
   806
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
   807
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   808
       \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   809
       @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   810
       \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   811
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   812
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   814
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   816
  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
   817
  as follows:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   818
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   819
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   820
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
   821
                          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
   822
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   823
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
   824
  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
   825
  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
   826
  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
   827
  @{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
   828
  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
   829
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   830
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   831
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
   832
apply(tactic {* select_tac' 4 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   833
apply(tactic {* select_tac' 3 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   834
apply(tactic {* select_tac' 2 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   835
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   836
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   837
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   838
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   839
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   840
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   841
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   842
  Since such repeated applications of a tactic to the reverse order of 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   843
  \emph{all} subgoals is quite common, there is the tactic combinator 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   844
  @{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
   845
  write: *}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   846
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   847
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
   848
apply(tactic {* ALLGOALS select_tac' *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   849
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   850
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   851
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   852
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   853
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   854
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   855
  Remember that we chose to implement @{ML select_tac'} so that it 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   856
  always  succeeds. This can be potentially very confusing for the user, 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   857
  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
   858
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   859
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   860
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   861
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   862
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   863
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   864
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   865
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   866
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   867
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   868
  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
   869
  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
   870
  therefore, tactics visible to the user should either change something or fail.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   871
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   872
  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
   873
  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
   874
  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
   875
  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
   876
  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
   877
  by the tactic. Because now
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   878
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   879
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   880
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   881
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   882
apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   883
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
   884
  \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
   885
  @{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
   886
  @{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
   887
  \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
   888
*}(*<*)oops(*>*)
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   889
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   890
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   891
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   892
  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
   893
  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
   894
  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
   895
  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
   896
*}
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
   897
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
   898
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
   899
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   900
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
   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
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
   903
apply(tactic {* repeat_xmp *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   904
txt{* produces
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   905
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   906
      \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
   907
      @{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
   908
      \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
   909
(*<*)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
   910
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
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
   912
  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
   913
  because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   914
  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
   915
  @{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
   916
  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
   917
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   918
  If you are after the ``primed'' version of @{ML repeat_xmp}, then you 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   919
  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
   920
*}
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
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   922
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
   923
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
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
   925
  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
   926
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   927
  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
   928
  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
   929
  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
   930
  is applied repeatedly only to the first subgoal. To analyse also all
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   931
  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
   932
  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
   933
*}
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
   934
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
   935
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
   936
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
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   938
  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
   939
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   940
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
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
   942
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
   943
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
   944
      @{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
   945
      \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
   946
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   947
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
   948
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   949
  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
   950
  include in @{ML select_tac'}. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   951
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   952
  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
   953
  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
   954
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
   955
*}
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
   956
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
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
   958
apply(tactic {* etac @{thm disjE} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   959
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
   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
      \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
   962
      @{subgoals [display]}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   963
      \end{minipage}\smallskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   964
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   965
      After typing
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   966
  *}
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
   967
(*<*)
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
   968
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
   969
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
   970
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
   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
back
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   973
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
   974
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
      \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
   976
      @{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
   977
      \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
   978
(*<*)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
   979
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   980
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
   981
  Sometimes this leads to confusing behaviour of tactics and also has
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   982
  the potential to explode the search space for tactics.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   983
  These problems can be avoided by prefixing the tactic with the tactic 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   984
  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
   985
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   986
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
   987
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
   988
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
   989
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
   990
      @{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
   991
      \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
   992
(*<*)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
   993
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   994
  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
   995
  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
   996
  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
   997
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   998
  \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
   999
  @{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
  1000
  @{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
  1001
  \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
  1002
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1003
  \begin{readmore}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1004
  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
  1005
  \end{readmore}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1006
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1007
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1008
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1009
section {* Simplifier Tactics *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1010
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1011
text {*
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1012
  A lot of convenience in the reasoning with Isabelle derives from its
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1013
  powerful simplifier. The power of simplifier is at the same time a
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1014
  strength and a weakness, because you can easily make the simplifier to
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1015
  loop and its behaviour can sometimes be difficult to predict. There is also
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1016
  a multitude of options that configurate the behaviour of the simplifier. We
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1017
  describe some of them in this and the next section.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1018
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1019
  There are the following five main tactics behind 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1020
  the simplifier (in parentheses is their user-level counterparts):
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1021
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1022
  \begin{isabelle}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1023
  \begin{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1024
  \begin{tabular}{l@ {\hspace{2cm}}l}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1025
   @{ML simp_tac}            & @{text "(simp (no_asm))"} \\
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1026
   @{ML asm_simp_tac}        & @{text "(simp (no_asm_simp))"} \\
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1027
   @{ML full_simp_tac}       & @{text "(simp (no_asm_use))"} \\
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1028
   @{ML asm_lr_simp_tac}     & @{text "(simp (asm_lr))"} \\
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1029
   @{ML asm_full_simp_tac}   & @{text "(simp)"}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1030
  \end{tabular}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1031
  \end{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1032
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1033
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1034
  All of them take a simpset and an interger as argument (the latter to specify the goal 
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1035
  to be analysed). So the proof
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1036
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1037
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1038
lemma "Suc (1 + 2) < 3 + 2"
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1039
apply(simp)
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1040
done
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1041
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1042
text {*
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1043
  corresponds on the ML-level to the tactic
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1044
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1045
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1046
lemma "Suc (1 + 2) < 3 + 2"
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1047
apply(tactic {* asm_full_simp_tac @{simpset} 1 *})
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1048
done
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1049
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1050
text {*
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1051
  If the simplifier cannot make any progress, then it leaves the goal unchanged
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1052
  and does not raise any error message. That means if you use it to unfold a definition
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1053
  for a constant and this constant is not present in a goal state, you can still
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1054
  safely apply the simplifier.
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1055
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1056
  When using the simplifier, the crucial information you have to control is
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1057
  the simpset. If not handled with care, then the simplifier can easily run
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1058
  into a loop. It might be surprising that a simpset is more complex than just a
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1059
  simple collection of theorems used as simplification rules. One reason for
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1060
  the complexity is that the simplifier must be able to rewrite inside terms
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1061
  and should also be able to rewrite according to rules that have
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1062
  precoditions.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1063
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1064
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1065
  The rewriting inside terms requires congruence rules, which
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1066
  are meta-equalities typical of the form
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1067
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1068
  \begin{isabelle}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1069
  \begin{center}
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1070
  \mbox{\inferrule{@{text "t\<^isub>i \<equiv> s\<^isub>i"}}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1071
                  {@{text "constr t\<^isub>1\<dots>t\<^isub>n \<equiv> constr s\<^isub>1\<dots>s\<^isub>n"}}}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1072
  \end{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1073
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1074
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1075
  with @{text "constr"} being a term-constructor. Every simpset contains only
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1076
  one concruence rule for each term-constructor, which however can be
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1077
  overwritten. The user can declare lemmas to be congruence rules using the
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1078
  attribute @{text "[cong]"}. In HOL, the user usually states these lemmas as
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1079
  equations, which are then internally transformed into meta-equations.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1080
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1081
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1082
  The rewriting with rules involving preconditions requires what is in
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1083
  Isabelle called a subgoaler, a solver and a looper. These can be arbitrary
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1084
  tactics that can be installed in a simpset. However, simpsets also include
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1085
  simprocs, which can produce rewrite rules on demand (see next
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1086
  section). Another component are split-rules, which can simplify for example
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1087
  the ``then'' and ``else'' branches of if-statements under the corresponding
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1088
  precoditions.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1089
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1090
  \begin{readmore}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1091
  For more information about the simplifier see @{ML_file "Pure/meta_simplifier.ML"}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1092
  and @{ML_file "Pure/simplifier.ML"}. The simplifier for HOL is set up in 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1093
  @{ML_file "HOL/Tools/simpdata.ML"}. Generic splitters are implemented in 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1094
  @{ML_file  "Provers/splitter.ML"}.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1095
  \end{readmore}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1096
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1097
  \begin{readmore}
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1098
  FIXME: Find the right place Discrimination nets are implemented
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1099
  in @{ML_file "Pure/net.ML"}.
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1100
  \end{readmore}
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1101
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1102
  The most common combinators to modify simpsets are
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1103
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1104
  \begin{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1105
  \begin{tabular}{ll}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1106
  @{ML addsimps}   & @{ML delsimps}\\
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1107
  @{ML addcongs}   & @{ML delcongs}\\
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1108
  @{ML addsimprocs} & @{ML delsimprocs}\\
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1109
  \end{tabular}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1110
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1111
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1112
  (FIXME: What about splitters? @{ML addsplits}, @{ML delsplits})
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1113
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1114
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1115
text_raw {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1116
\begin{figure}[tp]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1117
\begin{isabelle}*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1118
ML{*fun get_parts ss = 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1119
let 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1120
  val ({rules, ...}, {congs, procs, ...}) = MetaSimplifier.rep_ss ss
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1121
  
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1122
  val rules_list = Net.entries rules
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1123
  val rnames = map #name rules_list
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1124
  val rthms =  map #thm rules_list
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1125
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1126
  val congs_list = fst congs
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1127
  val cnames = map fst congs_list  
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1128
  val cthms = map (#thm o snd) congs_list
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1129
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1130
  val proc_list = Net.entries procs
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1131
in
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1132
  (rnames ~~ rthms, cnames ~~ cthms)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1133
end
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1134
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1135
fun print_parts ctxt ss =
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1136
let
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1137
  val (simps, congs) = get_parts ss
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1138
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1139
  fun name_thm (nm, thm) =
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1140
      "  " ^ nm ^ ": " ^ (str_of_thm ctxt thm)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1141
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1142
  val s1 = ["Simplification rules:"]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1143
  val s2 = map name_thm simps
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1144
  val s3 = ["Congruences rules:"]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1145
  val s4 = map name_thm congs
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1146
in
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1147
  (s1 @ s2 @ s3 @ s4) 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1148
     |> separate "\n"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1149
     |> implode
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1150
     |> warning
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1151
end*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1152
text_raw {* 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1153
\end{isabelle}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1154
\caption{The function @{ML MetaSimplifier.rep_ss} returns a record containing
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1155
  a simpset. We are here only interested in the simplifcation rules, congruence rules and
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1156
  simprocs. The first and third are discrimination nets, which from which we extract
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1157
  lists using @{ML Net.entries}. The congruence rules are stored in 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1158
  an association list, associating a constant with a rule.\label{fig:prettyss}}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1159
\end{figure} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1160
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1161
text {* 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1162
  To see how they work, consider the two functions in Figure~\ref{fig:prettyss}, which
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1163
  print out some parts of a simpset. If you use them to print out the components of the
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1164
  empty simpset, the most primitive simpset
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1165
  
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1166
  @{ML_response_fake [display,gray]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1167
  "print_parts @{context} empty_ss"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1168
"Simplification rules:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1169
Congruences rules:"}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1170
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1171
  you can see it contains nothing. This simpset is usually not useful, except as a 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1172
  building block to build bigger simpsets. For example you can add to @{ML empty_ss} 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1173
  the simplification rule @{thm [source] Diff_Int} as follows:
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1174
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1175
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1176
ML{*val ss1 = empty_ss addsimps [@{thm Diff_Int} RS @{thm eq_reflection}] *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1177
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1178
text {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1179
  Printing out the components of this simpset gives:
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1180
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1181
  @{ML_response_fake [display,gray]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1182
  "print_parts @{context} ss1"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1183
"Simplification rules:
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1184
  ??.unknown: A - B \<inter> C \<equiv> A - B \<union> (A - C)
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1185
Congruences rules:"}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1186
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1187
  (FIXME: Why does it print out ??.unknown)
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1188
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1189
  Adding the congruence rule @{thm [source] UN_cong} 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1190
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1191
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1192
ML{*val ss2 = ss1 addcongs [@{thm UN_cong} RS @{thm eq_reflection}] *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1193
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1194
text {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1195
  gives
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1196
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1197
  @{ML_response_fake [display,gray]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1198
  "print_parts @{context} ss2"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1199
"Simplification rules:
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1200
  ??.unknown: A - B \<inter> C \<equiv> A - B \<union> (A - C)
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1201
Congruences rules:
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1202
  UNION: \<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk> \<Longrightarrow> \<Union>x\<in>A. C x \<equiv> \<Union>x\<in>B. D x"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1203
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1204
  Notice that we had to add these lemmas as meta-equations. The @{ML empty_ss} 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1205
  expects this form of the simplification and congruence rules. However, even 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1206
  adding these lemmas to @{ML empty_ss} we do not end up with anything useful yet.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1207
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1208
  In the context of HOL the first useful simpset is @{ML HOL_basic_ss}. While
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1209
  printing out the components of this simpset
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1210
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1211
  @{ML_response_fake [display,gray]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1212
  "print_parts @{context} HOL_basic_ss"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1213
"Simplification rules:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1214
Congruences rules:"}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1215
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1216
  also produces ``nothing'', the printout is misleading. In fact
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1217
  the simpset @{ML HOL_basic_ss} is setup so that it can be used to solve goals of the
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1218
  form  @{thm TrueI}, @{thm refl[no_vars]}, @{term "t \<equiv> t"} and @{thm FalseE[no_vars]}, 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1219
  and resolve with assumptions.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1220
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1221
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1222
lemma 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1223
 "True" and "t = t" and "t \<equiv> t" and "False \<Longrightarrow> Foo" and "\<lbrakk>A; B; C\<rbrakk> \<Longrightarrow> A"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1224
apply(tactic {* ALLGOALS (simp_tac HOL_basic_ss) *})
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1225
done
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1226
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1227
text {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1228
  This is because how the tactics for the subgoaler, solver and looper are set 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1229
  up. 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1230
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1231
  The simpset @{ML HOL_ss}, which is an extention of @{ML HOL_basic_ss}, contains 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1232
  already many useful simplification rules for the logical connectives in HOL. 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1233
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1234
  @{ML_response_fake [display,gray]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1235
  "print_parts @{context} HOL_ss"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1236
"Simplification rules:
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1237
  Pure.triv_forall_equality: (\<And>x. PROP V) \<equiv> PROP V
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1238
  HOL.the_eq_trivial: THE x. x = y \<equiv> y
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1239
  HOL.the_sym_eq_trivial: THE ya. y = ya \<equiv> y
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1240
  \<dots>
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1241
Congruences rules:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1242
  HOL.simp_implies: \<dots>
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1243
    \<Longrightarrow> (PROP P =simp=> PROP Q) \<equiv> (PROP P' =simp=> PROP Q')
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1244
  op -->: \<lbrakk>P \<equiv> P'; P' \<Longrightarrow> Q \<equiv> Q'\<rbrakk> \<Longrightarrow> P \<longrightarrow> Q \<equiv> P' \<longrightarrow> Q'"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1245
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1246
  
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1247
  The main point of these simpsets is that they are small enough to
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1248
  not cause any loops with most of the simplification rules that
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1249
  you might like to add. 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1250
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1251
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1252
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1253
text_raw {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1254
\begin{figure}[tp]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1255
\begin{isabelle} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1256
types  prm = "(nat \<times> nat) list"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1257
consts perm :: "prm \<Rightarrow> 'a \<Rightarrow> 'a"  ("_ \<bullet> _" [80,80] 80)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1258
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1259
primrec (perm_nat)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1260
 "[]\<bullet>c = c"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1261
 "(ab#pi)\<bullet>c = (if (pi\<bullet>c)=fst ab 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1262
                 then snd ab 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1263
                 else (if (pi\<bullet>c)=snd ab then fst ab else (pi\<bullet>c)))" 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1264
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1265
primrec (perm_prod)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1266
 "pi\<bullet>(x, y) = (pi\<bullet>x, pi\<bullet>y)"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1267
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1268
primrec (perm_list)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1269
 "pi\<bullet>[] = []"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1270
 "pi\<bullet>(x#xs) = (pi\<bullet>x)#(pi\<bullet>xs)"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1271
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1272
lemma perm_append[simp]:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1273
  fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1274
  shows "((pi\<^isub>1@pi\<^isub>2)\<bullet>c) = (pi\<^isub>1\<bullet>(pi\<^isub>2\<bullet>c))"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1275
by (induct pi\<^isub>1) (auto) 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1276
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1277
lemma perm_eq[simp]:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1278
  fixes c::"nat" and pi::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1279
  shows "(pi\<bullet>c = pi\<bullet>d) = (c = d)" 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1280
by (induct pi) (auto)
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1281
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1282
lemma perm_rev[simp]:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1283
  fixes c::"nat" and pi::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1284
  shows "pi\<bullet>((rev pi)\<bullet>c) = c"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1285
by (induct pi arbitrary: c) (auto)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1286
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1287
lemma perm_compose:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1288
  fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1289
  shows "pi\<^isub>1\<bullet>(pi\<^isub>2\<bullet>c) = (pi\<^isub>1\<bullet>pi\<^isub>2)\<bullet>(pi\<^isub>1\<bullet>c)" 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1290
by (induct pi\<^isub>2) (auto)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1291
text_raw {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1292
\end{isabelle}\medskip
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1293
\caption{A simple theory about permutations over @{typ nat}. The point is that the
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1294
  lemma @{thm [source] perm_compose} cannot be directly added to the simplifier, as
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1295
  it would cause the simplifier to loop. It can still be used as a simplification 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1296
  rule if the permutation is sufficiently protected.\label{fig:perms}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1297
  (FIXME: Uses old primrec.)}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1298
\end{figure} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1299
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1300
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1301
text {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1302
  Often the situation arises that simplification rules will cause the
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1303
  simplifier to run into an infinite loop. Consider for example the simple
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1304
  theory about permutations shown in Figure~\ref{fig:perms}. The purpose of
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1305
  the lemmas is to pus permutations as far inside a term where they might
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1306
  disappear using the lemma @{thm [source] perm_rev}. However, to fully
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1307
  simplifiy all instances, it would be desirable to add also the lemma @{thm
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1308
  [source] perm_compose} to the simplifier in order to push permutations over
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1309
  other permutations. Unfortunately, the right-hand side of this lemma is
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1310
  again an instance of the left-hand side and so causes an infinite
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1311
  loop. On the user-level it is difficult to work around such situations
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1312
  and we end up with clunky proofs such as:
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1313
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1314
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1315
lemma 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1316
  fixes c d::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1317
  shows "pi\<^isub>1\<bullet>(c, pi\<^isub>2\<bullet>((rev pi\<^isub>1)\<bullet>d)) = (pi\<^isub>1\<bullet>c, (pi\<^isub>1\<bullet>pi\<^isub>2)\<bullet>d)"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1318
apply(simp)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1319
apply(rule trans)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1320
apply(rule perm_compose)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1321
apply(simp)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1322
done 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1323
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1324
text {*
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1325
  On the ML-level, we can however generate a single simplifier tactic that solves
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1326
  such proofs. The trick is to introduce an auxiliary constant for permutations 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1327
  and split the simplification into two phases. Let us say the auxiliary constant is
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1328
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1329
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1330
definition
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1331
  perm_aux :: "prm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet>aux _" [80,80] 80)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1332
where
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1333
  "pi \<bullet>aux c \<equiv> pi \<bullet> c"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1334
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1335
text {* Now the lemmas *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1336
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1337
lemma perm_aux_expand:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1338
  fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1339
  shows "pi\<^isub>1\<bullet>(pi\<^isub>2\<bullet>c) = pi\<^isub>1 \<bullet>aux (pi\<^isub>2\<bullet>c)" 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1340
unfolding perm_aux_def by (rule refl)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1341
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1342
lemma perm_compose_aux:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1343
  fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1344
  shows "pi\<^isub>1\<bullet>(pi\<^isub>2\<bullet>aux c) = (pi\<^isub>1\<bullet>pi\<^isub>2) \<bullet>aux (pi\<^isub>1\<bullet>c)" 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1345
unfolding perm_aux_def by (rule perm_compose)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1346
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1347
text {* 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1348
  are simple consequence of the definition and @{thm [source] perm_compose}. 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1349
  More importantly, the lemma @{thm [source] perm_compose_aux} can be safely 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1350
  added to the simplifier, because now the right-hand side is not anymore an instance 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1351
  of the left-hand side. So you can write 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1352
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1353
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1354
ML %linenosgray{*val perm_simp_tac =
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1355
let
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1356
  val thms1 = [@{thm perm_aux_expand}]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1357
  val thms2 = [@{thm perm_append}, @{thm perm_eq}, @{thm perm_rev}, 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1358
               @{thm perm_compose_aux}] @ @{thms perm_prod.simps} @ 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1359
               @{thms perm_list.simps} @ @{thms perm_nat.simps}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1360
  val thms3 = [@{thm perm_aux_def}]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1361
in
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1362
  simp_tac (HOL_basic_ss addsimps thms1)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1363
  THEN' simp_tac (HOL_basic_ss addsimps thms2)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1364
  THEN' simp_tac (HOL_basic_ss addsimps thms3)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1365
end*}
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1366
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1367
text {*
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1368
  This trick is not noticable for the user.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1369
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1370
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1371
lemma 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1372
  fixes c d::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1373
  shows "pi\<^isub>1\<bullet>(c, pi\<^isub>2\<bullet>((rev pi\<^isub>1)\<bullet>d)) = (pi\<^isub>1\<bullet>c, (pi\<^isub>1\<bullet>pi\<^isub>2)\<bullet>d)"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1374
apply(tactic {* perm_simp_tac 1 *})
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1375
done
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1376
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1377
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1378
text {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1379
  (FIXME: Is it interesting to say something about @{term "op =simp=>"}?)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1380
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1381
  (FIXME: Unfortunately one cannot access any simproc, as there is 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1382
  no @{text rep_proc} in function @{ML get_parts}.)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1383
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1384
  (FIXME: What are the second components of the congruence rules---something to
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1385
  do with weak congruence constants?)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1386
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1387
  (FIXME: Anything interesting to say about @{ML Simplifier.clear_ss}?)
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1388
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1389
  (FIXME: @{ML rewrite_goals_tac}, @{ML ObjectLogic.full_atomize_tac}, 
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1390
  @{ML ObjectLogic.rulify_tac})
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1391
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1392
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1393
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1394
section {* Simprocs *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1395
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1396
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1397
  In Isabelle you can also implement custom simplification procedures, called
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1398
  \emph{simprocs}. Simprocs can be triggered by the simplifier on a specified
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1399
  term-pattern and rewrite a term according to a theorem. They are useful in
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1400
  cases where a rewriting rule must be produced on ``demand'' or when
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1401
  rewriting by simplification is too unpredictable and potentially loops.
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1402
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1403
  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
  1404
  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
  1405
  you can use the function:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1406
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1407
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1408
ML %linenosgray{*fun fail_sp_aux simpset redex = 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1409
let
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1410
  val ctxt = Simplifier.the_context simpset
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1411
  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
  1412
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1413
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1414
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1415
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1416
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1417
  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
  1418
  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
  1419
  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
  1420
  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
  1421
  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
  1422
  that the simproc is triggered by the pattern @{term "Suc n"}. This can be
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1423
  done by adding the simproc to the current simpset as follows
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1424
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1425
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1426
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
  1427
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1428
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1429
  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
  1430
  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
  1431
  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
  1432
  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
  1433
  it fires on the lemma:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1434
*}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1435
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1436
lemma shows "Suc 0 = 1"
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1437
  apply(simp)
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1438
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1439
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1440
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1441
  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
  1442
  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
  1443
  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
  1444
  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
  1445
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1446
  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
  1447
  \isacommand{declare}-statement. For example the simproc will be deleted
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1448
  with the declaration
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1449
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1450
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1451
declare [[simproc del: fail_sp]]
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1452
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1453
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1454
  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
  1455
  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
  1456
  as follows:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1457
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1458
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1459
lemma shows "Suc 0 = 1"
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1460
  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
  1461
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1462
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1463
text {*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1464
  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
  1465
  left unchanged. 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1466
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1467
  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
  1468
  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
  1469
  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
  1470
  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
  1471
  to
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1472
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1473
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1474
ML{*fun fail_sp_aux' simpset redex = 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1475
let
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1476
  val ctxt = Simplifier.the_context simpset
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1477
  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
  1478
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1479
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1480
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1481
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1482
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1483
  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
  1484
  (therefore we printing it out using the function @{ML string_of_term in Syntax}).
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1485
  We can turn this function into a proper simproc using the function 
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1486
  @{ML Simplifier.simproc_i}:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1487
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1488
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1489
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1490
ML{*val fail_sp' = 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1491
let 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1492
  val thy = @{theory}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1493
  val pat = [@{term "Suc n"}]
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1494
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1495
  Simplifier.simproc_i thy "fail_sp'" pat (K fail_sp_aux')
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1496
end*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1497
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1498
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1499
  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
  1500
  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
  1501
  Now the simproc is set up and can be explicitly added using
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1502
  @{ML addsimprocs} to a simpset whenerver
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1503
  needed. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1504
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1505
  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
  1506
  see this in the proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1507
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1508
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1509
lemma shows "Suc (Suc 0) = (Suc 1)"
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1510
  apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [fail_sp']) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1511
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1512
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1513
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1514
  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
  1515
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1516
@{text [display]
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1517
"> Suc 0
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1518
> Suc (Suc 0) 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1519
> Suc 1"}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1520
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1521
  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
  1522
  rewrites terms according to the equation:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1523
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1524
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1525
lemma plus_one: 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1526
  shows "Suc n \<equiv> n + 1" by simp
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1527
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1528
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1529
  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
  1530
  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
  1531
  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
  1532
  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
  1533
  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
  1534
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1535
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1536
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1537
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
  1538
  case redex of
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1539
    @{term "Suc 0"} => NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1540
  | _ => SOME @{thm plus_one}*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1541
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1542
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1543
  and set up the simproc as follows.
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1544
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1545
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1546
ML{*val plus_one_sp =
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1547
let
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1548
  val thy = @{theory}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1549
  val pat = [@{term "Suc n"}] 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1550
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1551
  Simplifier.simproc_i thy "sproc +1" pat (K plus_one_sp_aux)
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1552
end*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1553
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1554
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1555
  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
  1556
  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
  1557
  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
  1558
  in the following proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1559
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1560
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1561
lemma shows "P (Suc (Suc (Suc 0))) (Suc 0)"
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1562
  apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [plus_one_sp]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1563
txt{*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1564
  where the simproc produces the goal state
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1565
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1566
  @{subgoals[display]}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1567
*}  
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1568
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1569
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1570
text {*
133
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
  1571
  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
  1572
  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
  1573
  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
  1574
  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
  1575
  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
  1576
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1577
  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
  1578
  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
  1579
  takes a term and produces the corresponding integer value.
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1580
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1581
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1582
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
  1583
  | 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
  1584
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1585
text {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1586
  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
  1587
  (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
  1588
  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
  1589
  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
  1590
  @{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
  1591
*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1592
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1593
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
  1594
let
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1595
  val num = HOLogic.mk_number @{typ "nat"} n
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1596
  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
  1597
in
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1598
  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
  1599
end*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1600
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1601
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1602
  From the integer value it generates the corresponding number term, called 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1603
  @{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
  1604
  (Line 4), which it proves by the arithmetic tactic in Line 6. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1605
134
328370b75c33 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
  1606
  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
  1607
  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
  1608
  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
  1609
  @{text "num_ss"} in the code will suffice.
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1610
*}
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1611
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1612
ML{*fun get_thm_alt ctxt (t, n) =
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1613
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1614
  val num = HOLogic.mk_number @{typ "nat"} n
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1615
  val goal = Logic.mk_equals (t, num)
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1616
  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
  1617
          @{thms nat_number} @ @{thms neg_simps} @ @{thms plus_nat.simps}
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1618
in
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1619
  Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1))
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1620
end*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1621
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1622
text {*
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1623
  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
  1624
  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
  1625
  what happens with @{ML arith_tac}, especially in more complicated 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1626
  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
  1627
  that is sufficiently powerful to solve every instance of the lemmas
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1628
  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
  1629
  ``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
  1630
  way than tracing the simplifier to obtain the lemmas that are successfully 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1631
  applied during simplification. Alas, there is none.} 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1632
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1633
  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
  1634
  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
  1635
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1636
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1637
ML{*fun nat_number_sp_aux ss t =
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1638
let 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1639
  val ctxt = Simplifier.the_context ss
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1640
in
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1641
  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
  1642
  handle TERM _ => NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1643
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1644
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1645
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1646
  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
  1647
  @{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
  1648
  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
  1649
  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
  1650
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1651
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1652
ML{*val nat_number_sp =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1653
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1654
  val thy = @{theory}
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1655
  val pat = [@{term "Suc n"}]
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1656
in 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1657
  Simplifier.simproc_i thy "nat_number" pat (K nat_number_sp_aux) 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1658
end*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1659
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1660
text {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1661
  Now in the lemma
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1662
  *}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1663
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1664
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
  1665
  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
  1666
txt {* 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1667
  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
  1668
  
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1669
  @{subgoals [display]}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1670
  *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1671
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1672
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1673
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1674
  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
  1675
  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
  1676
  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
  1677
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1678
  \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
  1679
  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
  1680
  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
  1681
  @{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
  1682
  \end{exercise}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1683
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1684
  (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
  1685
  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
  1686
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1687
137
a9685909944d new pfd file
Christian Urban <urbanc@in.tum.de>
parents: 135
diff changeset
  1688
section {* Conversions\label{sec:conversion} *}
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1689
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1690
text {*
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1691
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1692
  Conversions are a thin layer on top of Isabelle's inference kernel, and 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1693
  can be viewed be as a controllable, bare-bone version of Isabelle's simplifier.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1694
  One difference between conversions and the simplifier is that the former
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1695
  act on @{ML_type cterm}s while the latter acts on @{ML_type thm}s. 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1696
  However, we will also show in this section how conversions can be applied
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1697
  to theorems via tactics. The type for conversions is
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1698
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1699
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1700
ML{*type conv = Thm.cterm -> Thm.thm*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1701
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1702
text {*
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1703
  whereby the produced theorem is always a meta-equality. A simple conversion
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1704
  is the function @{ML "Conv.all_conv"}, which maps a @{ML_type cterm} to an
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1705
  instance of the (meta)reflexivity theorem. For example:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1706
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1707
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1708
  "Conv.all_conv @{cterm \"Foo \<or> Bar\"}"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1709
  "Foo \<or> Bar \<equiv> Foo \<or> Bar"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1710
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1711
  Another simple conversion is @{ML Conv.no_conv} which always raises the 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1712
  exception @{ML CTERM}.
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1713
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1714
  @{ML_response_fake [display,gray]
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1715
  "Conv.no_conv @{cterm True}" 
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1716
  "*** Exception- CTERM (\"no conversion\", []) raised"}
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1717
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1718
  A more interesting conversion is the function @{ML "Thm.beta_conversion"}: it 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1719
  produces a meta-equation between a term and its beta-normal form. For example
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1720
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1721
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1722
  "let
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1723
  val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1724
  val two = @{cterm \"2::nat\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1725
  val ten = @{cterm \"10::nat\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1726
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1727
  Thm.beta_conversion true (Thm.capply (Thm.capply add two) ten)
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1728
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1729
  "((\<lambda>x y. x + y) 2) 10 \<equiv> 2 + 10"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1730
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1731
  Note that the actual response in this example is @{term "2 + 10 \<equiv> 2 + (10::nat)"}, 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1732
  since the pretty-printer for @{ML_type cterm}s already beta-normalises terms.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1733
  But by the way how we constructed the term (using the function 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1734
  @{ML Thm.capply}, which is the application @{ML $} for @{ML_type cterm}s),
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1735
  we can be sure the left-hand side must contain beta-redexes. Indeed
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1736
  if we obtain the ``raw'' representation of the produced theorem, we
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1737
  can see the difference:
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1738
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1739
  @{ML_response [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1740
"let
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1741
  val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1742
  val two = @{cterm \"2::nat\"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1743
  val ten = @{cterm \"10::nat\"}
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1744
  val thm = Thm.beta_conversion true (Thm.capply (Thm.capply add two) ten)
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1745
in
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1746
  #prop (rep_thm thm)
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1747
end"
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1748
"Const (\"==\",\<dots>) $ 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1749
  (Abs (\"x\",\<dots>,Abs (\"y\",\<dots>,\<dots>)) $\<dots>$\<dots>) $ 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1750
    (Const (\"HOL.plus_class.plus\",\<dots>) $ \<dots> $ \<dots>)"} 
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1751
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1752
  The argument @{ML true} in @{ML Thm.beta_conversion} indicates that 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1753
  the right-hand side will be fully beta-normalised. If instead 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1754
  @{ML false} is given, then only a single beta-reduction is performed 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1755
  on the outer-most level. For example
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1756
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1757
  @{ML_response_fake [display,gray]
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1758
  "let
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1759
  val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1760
  val two = @{cterm \"2::nat\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1761
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1762
  Thm.beta_conversion false (Thm.capply add two)
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1763
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1764
  "((\<lambda>x y. x + y) 2) \<equiv> \<lambda>y. 2 + y"} 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1765
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1766
  Again, we actually see as output only the fully normalised term 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1767
  @{text "\<lambda>y. 2 + y"}.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1768
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1769
  The main point of conversions is that they can be used for rewriting
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1770
  @{ML_type cterm}s. To do this you can use the function @{ML
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1771
  "Conv.rewr_conv"}, which expects a meta-equation as an argument. Suppose we
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1772
  want to rewrite a @{ML_type cterm} according to the meta-equation:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1773
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1774
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1775
lemma true_conj1: "True \<and> P \<equiv> P" by simp
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1776
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1777
text {* 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1778
  You can see how this function works in the example rewriting 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1779
  the @{ML_type cterm} @{term "True \<and> (Foo \<longrightarrow> Bar)"} to @{term "Foo \<longrightarrow> Bar"}.
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1780
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1781
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1782
"let 
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1783
  val ctrm = @{cterm \"True \<and> (Foo \<longrightarrow> Bar)\"}
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1784
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1785
  Conv.rewr_conv @{thm true_conj1} ctrm
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1786
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1787
  "True \<and> (Foo \<longrightarrow> Bar) \<equiv> Foo \<longrightarrow> Bar"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1788
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1789
  Note, however, that the function @{ML Conv.rewr_conv} only rewrites the 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1790
  outer-most level of the @{ML_type cterm}. If the given @{ML_type cterm} does not match 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1791
  exactly the 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1792
  left-hand side of the theorem, then  @{ML Conv.rewr_conv} raises 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1793
  the exception @{ML CTERM}.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1794
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1795
  This very primitive way of rewriting can be made more powerful by
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1796
  combining several conversions into one. For this you can use conversion
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1797
  combinators. The simplest conversion combinator is @{ML then_conv}, 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1798
  which applies one conversion after another. For example
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1799
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1800
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1801
"let
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1802
  val conv1 = Thm.beta_conversion false
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1803
  val conv2 = Conv.rewr_conv @{thm true_conj1}
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1804
  val ctrm = Thm.capply @{cterm \"\<lambda>x. x \<and> False\"} @{cterm \"True\"}
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1805
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1806
  (conv1 then_conv conv2) ctrm
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1807
end"
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1808
  "(\<lambda>x. x \<and> False) True \<equiv> False"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1809
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1810
  where we first beta-reduce the term and then rewrite according to
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1811
  @{thm [source] true_conj1}. (Recall the problem with the pretty-printer
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1812
  normalising all terms.)
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1813
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1814
  The conversion combinator @{ML else_conv} tries out the 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1815
  first one, and if it does not apply, tries the second. For example 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1816
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1817
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1818
"let
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1819
  val conv = Conv.rewr_conv @{thm true_conj1} else_conv Conv.all_conv
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1820
  val ctrm1 = @{cterm \"True \<and> Q\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1821
  val ctrm2 = @{cterm \"P \<or> (True \<and> Q)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1822
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1823
  (conv ctrm1, conv ctrm2)
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1824
end"
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1825
"(True \<and> Q \<equiv> Q, P \<or> True \<and> Q \<equiv> P \<or> True \<and> Q)"}
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1826
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1827
  Here the conversion of @{thm [source] true_conj1} only applies
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1828
  in the first case, but fails in the second. The whole conversion
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  1829
  does not fail, however, because the combinator @{ML Conv.else_conv} will then 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1830
  try out @{ML Conv.all_conv}, which always succeeds.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1831
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1832
  Apart from the function @{ML beta_conversion in Thm}, which is able to fully
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1833
  beta-normalise a term, the conversions so far are restricted in that they
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1834
  only apply to the outer-most level of a @{ML_type cterm}. In what follows we
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1835
  will lift this restriction. The combinator @{ML Conv.arg_conv} will apply
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1836
  the conversion to the first argument of an application, that is the term
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1837
  must be of the form @{ML "t1 $ t2" for t1 t2} and the conversion is applied
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1838
  to @{text t2}.  For example
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1839
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1840
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1841
"let 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1842
  val conv = Conv.rewr_conv @{thm true_conj1}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1843
  val ctrm = @{cterm \"P \<or> (True \<and> Q)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1844
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1845
  Conv.arg_conv conv ctrm
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1846
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1847
"P \<or> (True \<and> Q) \<equiv> P \<or> Q"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1848
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1849
  The reason for this behaviour is that @{text "(op \<or>)"} expects two
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1850
  arguments.  Therefore the term must be of the form @{text "(Const \<dots> $ t1) $ t2"}. The
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1851
  conversion is then applied to @{text "t2"} which in the example above
150
cb39c41548bd added a comment about Conv.fun_conv
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
  1852
  stands for @{term "True \<and> Q"}. The function @{ML Conv.fun_conv} applies
cb39c41548bd added a comment about Conv.fun_conv
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
  1853
  the conversion to the first argument of an application.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1854
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1855
  The function @{ML Conv.abs_conv} applies a conversion under an abstractions.
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1856
  For example:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1857
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1858
  @{ML_response_fake [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1859
"let 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1860
  val conv = K (Conv.rewr_conv @{thm true_conj1}) 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1861
  val ctrm = @{cterm \"\<lambda>P. True \<and> P \<and> Foo\"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1862
in
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1863
  Conv.abs_conv conv @{context} ctrm
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1864
end"
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1865
  "\<lambda>P. True \<and> P \<and> Foo \<equiv> \<lambda>P. P \<and> Foo"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1866
  
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1867
  Note that this conversion needs a context as an argument. The conversion that 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1868
  goes under an application is @{ML Conv.combination_conv}. It expects two conversions 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1869
  as arguments, each of which is applied to the corresponding ``branch'' of the
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1870
  application. 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1871
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1872
  We can now apply all these functions in a conversion that recursively
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1873
  descends a term and applies a ``@{thm [source] true_conj1}''-conversion 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1874
  in all possible positions.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1875
*}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1876
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1877
ML %linenosgray{*fun all_true1_conv ctxt ctrm =
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1878
  case (Thm.term_of ctrm) of
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1879
    @{term "op \<and>"} $ @{term True} $ _ => 
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1880
      (Conv.arg_conv (all_true1_conv ctxt) then_conv
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1881
         Conv.rewr_conv @{thm true_conj1}) ctrm
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1882
  | _ $ _ => Conv.combination_conv 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1883
               (all_true1_conv ctxt) (all_true1_conv ctxt) ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1884
  | Abs _ => Conv.abs_conv (fn (_, ctxt) => all_true1_conv ctxt) ctxt ctrm
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1885
  | _ => Conv.all_conv ctrm*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1886
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1887
text {*
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1888
  This function ``fires'' if the terms is of the form @{text "True \<and> \<dots>"}; 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1889
  it descends under applications (Line 6 and 7) and abstractions 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1890
  (Line 8); otherwise it leaves the term unchanged (Line 9). In Line 2
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1891
  we need to transform the @{ML_type cterm} into a @{ML_type term} in order
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1892
  to be able to pattern-match the term. To see this 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1893
  conversion in action, consider the following example:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1894
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1895
@{ML_response_fake [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1896
"let
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1897
  val ctxt = @{context}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1898
  val ctrm = @{cterm \"distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x\"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1899
in
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1900
  all_true1_conv ctxt ctrm
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1901
end"
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1902
  "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
  1903
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1904
  To see how much control you have about rewriting by using conversions, let us 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1905
  make the task a bit more complicated by rewriting according to the rule
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1906
  @{text true_conj1}, but only in the first arguments of @{term If}s. Then 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1907
  the conversion should be as follows.
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1908
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1909
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1910
ML{*fun if_true1_conv ctxt ctrm =
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1911
  case Thm.term_of ctrm of
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1912
    Const (@{const_name If}, _) $ _ =>
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1913
      Conv.arg_conv (all_true1_conv ctxt) ctrm
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1914
  | _ $ _ => Conv.combination_conv 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1915
                (if_true1_conv ctxt) (if_true1_conv ctxt) ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1916
  | Abs _ => Conv.abs_conv (fn (_, ctxt) => if_true1_conv ctxt) ctxt ctrm
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1917
  | _ => Conv.all_conv ctrm *}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1918
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1919
text {*
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1920
  Here is an example for this conversion:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1921
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1922
  @{ML_response_fake [display,gray]
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1923
"let
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1924
  val ctxt = @{context}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1925
  val ctrm = 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1926
       @{cterm \"if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False\"}
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1927
in
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1928
  if_true1_conv ctxt ctrm
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1929
end"
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1930
"if P (True \<and> 1 \<noteq> 2) then True \<and> True else True \<and> False 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1931
\<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
  1932
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1933
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1934
text {*
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1935
  So far we only applied conversions to @{ML_type cterm}s. Conversions can, however, 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1936
  also work on theorems using the function @{ML "Conv.fconv_rule"}. As an example, 
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1937
  consider the conversion @{ML all_true1_conv} and the lemma:
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1938
*}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1939
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1940
lemma foo_test: "P \<or> (True \<and> \<not>P)" by simp
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1941
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1942
text {*
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1943
  Using the conversion you can transform this theorem into a new theorem
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1944
  as follows
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1945
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1946
  @{ML_response_fake [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1947
  "Conv.fconv_rule (all_true1_conv @{context}) @{thm foo_test}" 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1948
  "?P \<or> \<not> ?P"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1949
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1950
  Finally, conversions can also be turned into tactics and then applied to
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1951
  goal states. This can be done with the help of the function @{ML CONVERSION},
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1952
  and also some predefined conversion combinators that traverse a goal
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1953
  state. The combinators for the goal state are: @{ML Conv.params_conv} for
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1954
  converting under parameters (i.e.~where goals are of the form @{text "\<And>x. P \<Longrightarrow>
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1955
  Q"}); the function @{ML Conv.prems_conv} for applying a conversion to all
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1956
  premises of a goal, and @{ML Conv.concl_conv} for applying a conversion to
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1957
  the conclusion of a goal.
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1958
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1959
  Assume we want to apply @{ML all_true1_conv} only in the conclusion
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1960
  of the goal, and @{ML if_true1_conv} should only apply to the premises.
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1961
  Here is a tactic doing exactly that:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1962
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1963
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1964
ML{*val true1_tac = CSUBGOAL (fn (goal, i) =>
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1965
  let 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1966
    val ctxt = ProofContext.init (Thm.theory_of_cterm goal)
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1967
  in
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1968
    CONVERSION 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1969
      (Conv.params_conv ~1 (fn ctxt =>
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  1970
        (Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  1971
          Conv.concl_conv ~1 (all_true1_conv ctxt))) ctxt) i
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1972
  end)*}
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1973
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1974
text {* 
148
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  1975
  We call the conversions with the argument @{ML "~1"}. This is to 
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  1976
  analyse all parameters, premises and conclusions. If we call them with 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1977
  a non-negative number, say @{text n}, then these conversions will 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1978
  only be called on @{text n} premises (similar for parameters and
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1979
  conclusions). To test the tactic, consider the proof
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1980
*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1981
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1982
lemma
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1983
  "if True \<and> P then P else True \<and> False \<Longrightarrow>
148
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  1984
  (if True \<and> Q then True \<and> Q else P) \<longrightarrow> True \<and> (True \<and> Q)"
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1985
apply(tactic {* true1_tac 1 *})
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1986
txt {* where the tactic yields the goal state
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1987
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1988
       @{subgoals [display]}*}
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1989
(*<*)oops(*>*)
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1990
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1991
text {*
148
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  1992
  As you can see, the premises are rewritten according to @{ML if_true1_conv}, while
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  1993
  the conclusion according to @{ML all_true1_conv}.
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  1994
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1995
  To sum up this section, conversions are not as powerful as the simplifier
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1996
  and simprocs; the advantage of conversions, however, is that you never have
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1997
  to worry about non-termination.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1998
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  1999
  \begin{exercise}\label{ex:addconversion}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2000
  Write a tactic that does the same as the simproc in exercise
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2001
  \ref{ex:addsimproc}, but is based in conversions. That means replace terms
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2002
  of the form @{term "t\<^isub>1 + t\<^isub>2"} by their result. You can make
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2003
  the same assumptions as in \ref{ex:addsimproc}. FIXME: the current solution
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2004
  requires some additional functions.
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2005
  \end{exercise}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2006
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2007
  \begin{exercise}
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  2008
  Compare which way (either Exercise ~\ref{addsimproc} or \ref{ex:addconversion}) of 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  2009
  rewriting such terms is faster. For this you might have to construct quite 
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  2010
  large terms. Also see Recipe \ref{rec:timing} for information about timing.
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  2011
  \end{exercise}
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  2012
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2013
  \begin{readmore}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2014
  See @{ML_file "Pure/conv.ML"} for more information about conversion combinators. 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2015
  Further conversions are defined in  @{ML_file "Pure/thm.ML"},
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2016
  @{ML_file "Pure/drule.ML"} and @{ML_file "Pure/meta_simplifier.ML"}.
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2017
  \end{readmore}
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  2018
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2019
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2020
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2021
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2022
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
  2023
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2024
section {* Structured Proofs (TBD) *}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2025
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2026
text {* TBD *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2027
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2028
lemma True
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2029
proof
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2030
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2031
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2032
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2033
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2034
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2035
    then have "A & B" ..
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2036
    then have C by (rule r)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2037
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2038
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2039
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2040
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2041
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2042
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2043
    note conjI [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2044
    note r [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2045
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2046
oops
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2047
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2048
ML {* fun prop ctxt s =
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2049
  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
  2050
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2051
ML {* 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2052
  val ctxt0 = @{context};
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2053
  val ctxt = ctxt0;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2054
  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
  2055
  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
  2056
  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
  2057
  val this = [@{thm conjI} OF this]; 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2058
  val this = r OF this;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2059
  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
  2060
  val this = Variable.export ctxt ctxt0 [this] 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2061
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2062
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2063
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
  2064
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2065
end