ProgTutorial/Tactical.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 07 Oct 2009 11:28:40 +0200
changeset 335 163ac0662211
parent 334 4ae1ecb71539
child 339 c588e8422737
permissions -rw-r--r--
reorganised the certified terms section; tuned
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 {*
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
     8
  One of the main reason for descending to the ML-level of Isabelle is to be able to
105
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.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
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
  
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    52
  To start the proof, the function @{ML_ind "Goal.prove"}~@{text "ctxt xs As C
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    53
  tac"} sets up a goal state for proving the goal @{text C} (that is @{prop "P
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    54
  \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the assumptions @{text As}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    55
  (happens to be empty) with the variables @{text xs} that will be generalised
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    56
  once the goal is proved (in our case @{text P} and @{text
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    57
  Q}).\footnote{FIXME: explain prove earlier} The @{text "tac"} is the tactic
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    58
  that proves the goal; it can make use of the local assumptions (there are
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    59
  none in this example). The tactics @{ML_ind etac}, @{ML_ind rtac} and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    60
  @{ML_ind atac} in the code above correspond roughly to @{text erule}, @{text
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    61
  rule} and @{text assumption}, respectively. The operator @{ML_ind THEN}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    62
  strings the tactics together.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  \begin{readmore}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    65
  To learn more about the function @{ML_ind prove in Goal} see
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    66
  \isccite{sec:results} and the file @{ML_file "Pure/goal.ML"}.  See @{ML_file
289
08ffafe2585d adapted to changes in Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 288
diff changeset
    67
  "Pure/tactic.ML"} and @{ML_file "Pure/tactical.ML"} for the code of basic
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    68
  tactics and tactic combinators; see also Chapters 3 and 4 in the old
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    69
  Isabelle Reference Manual, and Chapter 3 in the Isabelle/Isar Implementation
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    70
  Manual.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    73
  Note that in the code above we use antiquotations for referencing the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    74
  theorems. Many theorems also have ML-bindings with the same name. Therefore,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    75
  we could also just have written @{ML "etac disjE 1"}, or in case where there
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    76
  is no ML-binding obtain the theorem dynamically using the function @{ML
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    77
  thm}; for example \mbox{@{ML "etac (thm \"disjE\") 1"}}. Both ways however
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    78
  are considered bad style! The reason is that the binding for @{ML disjE} can
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    79
  be re-assigned by the user and thus one does not have complete control over
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    80
  which theorem is actually applied. This problem is nicely prevented by using
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    81
  antiquotations, because then the theorems are fixed statically at
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    82
  compile-time.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    84
  During the development of automatic proof procedures, you will often find it
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    85
  necessary to test a tactic on examples. This can be conveniently done with
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    86
  the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  Consider the following sequence of tactics
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    88
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
ML{*val foo_tac = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
      (etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
       THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
       THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
       THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
       THEN atac 1)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
text {* and the Isabelle proof: *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
lemma "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
apply(tactic {* foo_tac *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   105
  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
   106
  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
   107
  any other function that returns a tactic.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
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
   109
  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
   110
  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
   111
  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
   112
  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
   113
  of goals is sometimes wanted, but usually it is not. To avoid the explicit numbering, 
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   114
  you can write
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
ML{*val foo_tac' = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
      (etac @{thm disjE} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
       THEN' rtac @{thm disjI2} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
       THEN' atac 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
       THEN' rtac @{thm disjI1} 
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   122
       THEN' atac)*}text_raw{*\label{tac:footacprime}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
text {* 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   125
  where @{ML_ind  THEN'} is used instead of @{ML THEN}. With @{ML foo_tac'} you can give 
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   126
  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
   127
  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
   128
  subsequently the first.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
   and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
apply(tactic {* foo_tac' 2 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
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
   138
  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
   139
  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
   140
  (@{ML THEN} is only one such operator) a ``primed'' version exists.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   141
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
   142
  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
   143
  analysing goals being only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
   144
  of this form, then these tactics return the error message:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   145
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   146
  \begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   147
  @{text "*** empty result sequence -- proof command failed"}\\
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   148
  @{text "*** At command \"apply\"."}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   149
  \end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   150
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   151
  This means they failed.\footnote{To be precise, tactics do not produce this error
288
griff
parents: 256
diff changeset
   152
  message, it originates from the \isacommand{apply} wrapper.} The reason for this 
241
29d4701c5ee1 polished
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   153
  error message is that tactics 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   154
  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
   155
  Hence the type of a tactic is:
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   156
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   158
ML{*type tactic = thm -> thm Seq.seq*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   160
text {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   161
  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
   162
  Therefore, if you write your own tactics, they  should not raise exceptions 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   163
  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
   164
  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
   165
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   166
  The simplest tactics are @{ML_ind  no_tac} and @{ML_ind  all_tac}. The first returns
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
   167
  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
   168
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   169
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   170
ML{*fun 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
   171
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   172
text {*
5e309df58557 general 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
  which means @{ML no_tac} always fails. The second returns the given theorem wrapped 
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   174
  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
   175
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   176
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   177
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
   178
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   179
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   180
  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
   181
  with the proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   183
  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
   184
  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
   185
  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
   186
  @{ML foo_tac'}: either using the first assumption or the second.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
back
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
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
   194
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   196
  By using \isacommand{back}, we construct the proof that uses the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   197
  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
   198
  assumption is used, in more interesting cases provability might depend
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   199
  on exploring different possibilities.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   200
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  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
   203
  sequences. In day-to-day Isabelle programming, however, one rarely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   204
  constructs sequences explicitly, but uses the predefined tactics and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   205
  tactic combinators instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   208
  It might be surprising that tactics, which transform
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   209
  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
   210
  (sequences). The surprise resolves by knowing that every 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   211
  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
   212
  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
   213
  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
   214
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   215
5e309df58557 general 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
ML{*fun my_print_tac ctxt thm =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   217
let
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   218
  val _ = tracing (string_of_thm_no_vars ctxt thm)
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   219
in 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   220
  Seq.single thm
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   221
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
   222
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   223
text_raw {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   224
  \begin{figure}[p]
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   225
  \begin{boxedminipage}{\textwidth}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   226
  \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
   227
*}
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   228
notation (output) "prop"  ("#_" [1000] 1000)
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   229
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   230
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
   231
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
   232
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   233
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
   234
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   235
      \end{minipage}\medskip      
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   236
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   237
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   238
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   239
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   240
      internal goal state:\\
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   241
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   242
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   243
      \end{minipage}\medskip
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
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
   246
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
   247
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
   248
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   249
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
   250
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   251
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   252
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   253
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   254
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   255
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   256
      internal goal state:\\
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   257
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   258
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   259
      \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
   260
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   261
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   262
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
   263
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
   264
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   265
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
   266
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   267
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   268
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   269
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   270
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   271
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   272
      internal goal state:\\
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   273
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   274
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   275
      \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
   276
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   277
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   278
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
   279
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
   280
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   281
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
   282
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   283
      \end{minipage}\medskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   284
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   285
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   286
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   287
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   288
      internal goal state:\\
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   289
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   290
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   291
      \end{minipage}\medskip   
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   292
   *}
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   293
(*<*)oops(*>*)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   294
text_raw {*  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   295
  \end{isabelle}
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   296
  \end{boxedminipage}
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   297
  \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
   298
  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
   299
  the goal state as represented internally (highlighted boxes). This
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   300
  tactic shows that every goal state in Isabelle is represented by a theorem:
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   301
  when you start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"}} the theorem is
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   302
  @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> #(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when you finish the proof the
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   303
  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
   304
  \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
   305
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   306
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   307
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   308
  which prints out the given theorem (using the string-function defined in
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   309
  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
   310
  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
   311
  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
   312
  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
   313
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   314
  @{text[display] "A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> #C"}
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
   315
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   316
  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
   317
  the subgoals. So after setting up the lemma, the goal state is always of the
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   318
  form @{text "C \<Longrightarrow> #C"}; when the proof is finished we are left with @{text
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
   319
  "#C"}. Since the goal @{term C} can potentially be an implication, there is a
241
29d4701c5ee1 polished
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   320
  ``protector'' wrapped around it (the wrapper is the outermost constant
303
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   321
  @{text "Const (\"prop\", bool \<Rightarrow> bool)"}; in the figure we make it visible
05e6a33edef6 added an antiquotation for printing the raw proof state; polished the example about proof state
Christian Urban <urbanc@in.tum.de>
parents: 302
diff changeset
   322
  as an @{text #}). This wrapper prevents that premises of @{text C} are
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
   323
  misinterpreted as open subgoals. While tactics can operate on the subgoals
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   324
  (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
   325
  @{term C} intact, with the exception of possibly instantiating schematic
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   326
  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
   327
  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
   328
 
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
   329
  \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
   330
  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
   331
  \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
   332
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
   333
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   334
194
8cd51a25a7ca some polishing
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   335
section {* Simple Tactics\label{sec:simpletacs} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   337
text {*
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   338
  Let us start with explaining the simple tactic @{ML_ind  print_tac}, which is quite useful 
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   339
  for low-level debugging of tactics. It just prints out a message and the current 
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   340
  goal state. Unlike @{ML my_print_tac} shown earlier, it prints the goal state 
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   341
  as the user would see it.  For example, processing the proof
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   342
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   343
 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   344
lemma shows "False \<Longrightarrow> True"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   345
apply(tactic {* print_tac "foo message" *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   346
txt{*gives:\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   347
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   348
     \begin{minipage}{\textwidth}\small
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   349
     @{text "foo message"}\\[3mm] 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   350
     @{prop "False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   351
     @{text " 1. False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   352
     \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   353
   *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   354
(*<*)oops(*>*)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   355
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   356
text {*
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   357
  A simple tactic for easy discharge of any proof obligations is 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   358
  @{ML_ind  cheat_tac in SkipProof}. This tactic corresponds to
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   359
  the Isabelle command \isacommand{sorry} and is sometimes useful  
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   360
  during the development of tactics.
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   361
*}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   362
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   363
lemma shows "False" and "Goldbach_conjecture"  
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   364
apply(tactic {* SkipProof.cheat_tac @{theory} *})
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   365
txt{*\begin{minipage}{\textwidth}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   366
     @{subgoals [display]}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   367
     \end{minipage}*}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   368
(*<*)oops(*>*)
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   369
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   370
text {*
241
29d4701c5ee1 polished
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   371
  This tactic works however only if the quick-and-dirty mode of Isabelle 
29d4701c5ee1 polished
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   372
  is switched on.
29d4701c5ee1 polished
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   373
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   374
  Another simple tactic is the function @{ML_ind  atac}, which, as shown in the previous
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   375
  section, corresponds to the assumption command.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   376
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   377
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   378
lemma shows "P \<Longrightarrow> P"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
apply(tactic {* atac 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   380
txt{*\begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   381
     @{subgoals [display]}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   382
     \end{minipage}*}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   383
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   385
text {*
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   386
  Similarly, @{ML_ind  rtac}, @{ML_ind  dtac}, @{ML_ind  etac} and 
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   387
  @{ML_ind  ftac} correspond (roughly)
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
   388
  to @{text rule}, @{text drule}, @{text erule} and @{text frule}, 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   389
  respectively. Each of them takes a theorem as argument and attempts to 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   390
  apply it to a goal. Below are three self-explanatory examples.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   391
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   392
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   393
lemma shows "P \<and> Q"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
apply(tactic {* rtac @{thm conjI} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   395
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   396
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   397
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   400
lemma shows "P \<and> Q \<Longrightarrow> False"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
apply(tactic {* etac @{thm conjE} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   402
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   403
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   404
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
lemma shows "False \<and> True \<Longrightarrow> False"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
apply(tactic {* dtac @{thm conjunct2} 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   409
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   410
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   411
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
text {*
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   415
  The function @{ML_ind  resolve_tac} is similar to @{ML_ind  rtac}, except that it
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   416
  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
   417
  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
   418
  explored via the lazy sequences mechanism). Given the code
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   421
ML{*val resolve_xmp_tac = resolve_tac [@{thm impI}, @{thm conjI}]*}
99
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 {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   424
  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
   425
  implication is analysed and then an outermost conjunction.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   426
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   427
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   428
lemma shows "C \<longrightarrow> (A \<and> B)" and "(A \<longrightarrow> B) \<and> C"
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   429
apply(tactic {* resolve_xmp_tac 1 *})
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   430
apply(tactic {* resolve_xmp_tac 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   431
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   432
     @{subgoals [display]} 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   433
     \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   434
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   435
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   436
text {* 
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
   437
  Similar versions taking a list of theorems exist for the tactics 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   438
  @{ML dtac} (@{ML_ind  dresolve_tac}), @{ML etac} 
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   439
  (@{ML_ind  eresolve_tac}) and so on.
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   440
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   441
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   442
  Another simple tactic is @{ML_ind  cut_facts_tac}. It inserts a list of theorems
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   443
  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
   444
*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   445
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   446
lemma shows "True \<noteq> False"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   447
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
   448
txt{*produces the goal state\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   449
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   450
     \begin{minipage}{\textwidth}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   451
     @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   452
     \end{minipage}*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   453
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   454
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   455
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   456
  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
   457
  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
   458
  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
   459
  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
   460
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   461
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   462
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
   463
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
   464
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
   465
     @{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
   466
     \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
   467
(*<*)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
   468
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   469
text {*
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   470
  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
   471
  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
   472
  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
   473
  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
   474
  should be, then this situation can be avoided by introducing a more
241
29d4701c5ee1 polished
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   475
  constrained version of the @{text bspec}-rule. Such constraints can be given by
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   476
  pre-instantiating theorems with other theorems. One function to do this is
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   477
  @{ML_ind "RS"}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   478
  
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   479
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   480
  "@{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
   481
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   482
  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
   483
  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
   484
  case of
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   485
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   486
  @{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
   487
  "@{thm conjI} RS @{thm mp}" 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   488
"*** 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
   489
[\"\<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
   490
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   491
  then the function raises an exception. The function @{ML_ind  RSN} is similar to @{ML RS}, but 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   492
  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
   493
  should be instantiated. 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   494
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   495
  To improve readability of the theorems we shall produce below, we will use the
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   496
  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
   497
  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
   498
  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
   499
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   500
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   501
  "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
   502
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   503
  If you want to instantiate more than one premise of a theorem, you can use 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   504
  the function @{ML_ind  MRS}:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   505
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   506
  @{ML_response_fake [display,gray]
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   507
  "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
   508
  "\<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
   509
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   510
  If you need to instantiate lists of theorems, you can use the
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   511
  functions @{ML RL} and @{ML_ind  MRL}. For example in the code below,
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   512
  every theorem in the second list is instantiated with every 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   513
  theorem in the first.
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   514
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   515
  @{ML_response_fake [display,gray]
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   516
"map (no_vars @{context}) 
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   517
   ([@{thm impI}, @{thm disjI2}] RL [@{thm conjI}, @{thm disjI1}])" 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   518
"[\<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
   519
 \<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
   520
 (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
   521
 Q \<Longrightarrow> (P \<or> Q) \<or> Qa]"}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   522
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   523
  \begin{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   524
  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
   525
  \end{readmore}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   526
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   527
  Often proofs on the ML-level involve elaborate operations on assumptions and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   528
  @{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
   529
  shown so far is very unwieldy and brittle. Some convenience and
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   530
  safety is provided by the functions @{ML_ind  FOCUS in Subgoal} and 
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   531
  @{ML_ind  SUBPROOF}. These tactics fix the parameters 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   532
  and bind the various components of a goal state to a record. 
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   533
  To see what happens, use the function defined in Figure~\ref{fig:sptac}, which
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   534
  takes a record and just prints out the contents of this record. Consider
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   535
  now the proof:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   536
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   537
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   538
text_raw{*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   539
\begin{figure}[t]
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
   540
\begin{minipage}{\textwidth}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   541
\begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   542
*}
294
ee9d53fbb56b made changes for SUBPROOF and sat_tac
Christian Urban <urbanc@in.tum.de>
parents: 292
diff changeset
   543
ee9d53fbb56b made changes for SUBPROOF and sat_tac
Christian Urban <urbanc@in.tum.de>
parents: 292
diff changeset
   544
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   545
ML{*fun foc_tac {prems, params, asms, concl, context, schematics} = 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   546
let 
294
ee9d53fbb56b made changes for SUBPROOF and sat_tac
Christian Urban <urbanc@in.tum.de>
parents: 292
diff changeset
   547
  val string_of_params = string_of_cterms context (map snd params)
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 243
diff changeset
   548
  val string_of_asms = string_of_cterms context asms
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 243
diff changeset
   549
  val string_of_concl = string_of_cterm context concl
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 243
diff changeset
   550
  val string_of_prems = string_of_thms_no_vars context prems   
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   551
  val string_of_schms = string_of_cterms context (map fst (snd schematics))    
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   552
 
305
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   553
  val strs = ["params: " ^ string_of_params,
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   554
              "schematics: " ^ string_of_schms,
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   555
              "assumptions: " ^ string_of_asms,
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   556
              "conclusion: " ^ string_of_concl,
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   557
              "premises: " ^ string_of_prems]
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   558
in
305
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 303
diff changeset
   559
  tracing (cat_lines strs); all_tac 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   560
end*}
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   561
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   562
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   563
\end{isabelle}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
   564
\end{minipage}
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   565
\caption{A function that prints out the various parameters provided by 
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   566
 @{ML FOCUS in Subgoal} and @{ML SUBPROOF}. It uses the functions defined 
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   567
  in Section~\ref{sec:printing} for extracting strings from @{ML_type cterm}s 
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   568
  and @{ML_type thm}s.\label{fig:sptac}}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   569
\end{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   570
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   571
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   572
lemma shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x"
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   573
apply(tactic {* Subgoal.FOCUS foc_tac @{context} 1 *})
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   574
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   575
txt {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   576
  The tactic produces the following printout:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   577
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   578
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   579
  \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
   580
  params:      & @{term x}, @{term y}\\
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   581
  schematics:  & @{text ?z}\\
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
   582
  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
   583
  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
   584
  premises:    & @{term "A x y"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   585
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   586
  \end{quote}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   587
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   588
  (FIXME: Find out how nowadays the schematics are handled)
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   589
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   590
  Notice in the actual output the brown colour of the variables @{term x}, 
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   591
  and @{term y}. Although they are parameters in the original goal, they are fixed inside
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   592
  the tactic. By convention these fixed variables are printed in brown colour.
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   593
  Similarly the schematic variable @{text ?z}. The assumption, or premise, 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   594
  @{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
   595
  @{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
   596
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   597
  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
   598
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   599
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   600
apply(rule impI)
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   601
apply(tactic {* Subgoal.FOCUS foc_tac @{context} 1 *})
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   602
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   603
txt {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   604
  then the tactic prints out: 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   605
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   606
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   607
  \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
   608
  params:      & @{term x}, @{term y}\\
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   609
  schematics:  & @{text ?z}\\
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
   610
  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
   611
  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
   612
  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
   613
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   614
  \end{quote}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   615
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   616
(*<*)oops(*>*)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   617
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   618
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   619
  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
   620
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   621
  The difference between the tactics @{ML SUBPROOF} and @{ML FOCUS in Subgoal}
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   622
  is that the former expects that the goal is solved completely, which the
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   623
  latter does not. @{ML SUBPROOF} can also not instantiate an schematic
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   624
  variables.
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   625
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   626
  One convenience of both @{ML FOCUS in Subgoal} and @{ML SUBPROOF} is that we
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   627
  can apply the assumptions using the usual tactics, because the parameter
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   628
  @{text prems} contains them as theorems. With this you can easily implement
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   629
  a tactic that behaves almost like @{ML atac}:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   630
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   631
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   632
ML{*val atac' = Subgoal.FOCUS (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
   633
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   634
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   635
  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
   636
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   637
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   638
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
   639
apply(tactic {* atac' @{context} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   640
txt{* it will produce
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   641
      @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   642
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   643
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   644
text {*
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   645
  Notice that @{ML atac'} inside @{ML FOCUS in Subgoal} calls @{ML
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   646
  resolve_tac} with the subgoal number @{text "1"} and also the outer call to
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   647
  @{ML FOCUS in Subgoal} in the \isacommand{apply}-step uses @{text "1"}. This
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   648
  is another advantage of @{ML FOCUS in Subgoal} and @{ML SUBPROOF}: the
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   649
  addressing inside it is completely local to the tactic inside the
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   650
  subproof. It is therefore possible to also apply @{ML atac'} to the second
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   651
  goal by just writing:
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   652
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   653
*}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   654
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   655
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
   656
apply(tactic {* atac' @{context} 2 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   657
apply(rule TrueI)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   658
done
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   659
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   660
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
text {*
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   662
  \begin{readmore}
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   663
  The functions @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are defined in 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   664
  @{ML_file "Pure/subgoal.ML"} and also described in 
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   665
  \isccite{sec:results}. 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   666
  \end{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   667
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   668
  Similar but less powerful functions than @{ML FOCUS in Subgoal}, respectively
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   669
  @{ML SUBPROOF}, are 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   670
  @{ML_ind  SUBGOAL} and @{ML_ind  CSUBGOAL}. They allow you to 
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   671
  inspect a given subgoal (the former
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
   672
  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
   673
  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
   674
  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
   675
  analyse a few connectives). The code of the tactic is as
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   676
  follows.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
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
   679
ML %linenosgray{*fun select_tac (t, i) =
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   680
  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
   681
     @{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
   682
   | @{term "op \<Longrightarrow>"} $ _ $ t' => select_tac (t', i)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   683
   | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   684
   | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   685
   | @{term "Not"} $ _ => rtac @{thm notI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   686
   | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   687
   | _ => all_tac*}text_raw{*\label{tac:selecttac}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   688
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   689
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   690
  The input of the function is a term representing the subgoal and a number
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
   691
  specifying the subgoal of interest. In Line 3 you need to descend under the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   692
  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
   693
  analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   694
  analysed. Similarly with meta-implications in the next line.  While for the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   695
  first five patterns we can use the @{text "@term"}-antiquotation to
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   696
  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
   697
  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
   698
  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
   699
  basic term-constructors. This is not necessary in other cases, because their
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   700
  type is always fixed to function types involving only the type @{typ
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   701
  bool}. (See Section \ref{sec:terms_types_manually} about constructing terms
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   702
  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
   703
  Consequently, @{ML select_tac} never fails.
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   704
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   705
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   706
  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
   707
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   708
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   709
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   710
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
   711
apply(tactic {* SUBGOAL select_tac 4 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   712
apply(tactic {* SUBGOAL select_tac 3 *})
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   713
apply(tactic {* SUBGOAL select_tac 2 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   714
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
   715
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   716
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   717
      \end{minipage} *}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   718
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   719
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   720
text {*
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   721
  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
   722
  Note that we applied the tactic to the goals in ``reverse'' order. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   723
  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
   724
  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
   725
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   726
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   727
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
   728
apply(tactic {* SUBGOAL select_tac 1 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   729
apply(tactic {* SUBGOAL select_tac 3 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   730
apply(tactic {* SUBGOAL select_tac 4 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   731
apply(tactic {* SUBGOAL select_tac 5 *})
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   732
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   733
105
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
  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
   736
  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
   737
  the ``reverse application'' is easy to implement.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   738
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
   739
  Of course, this example is
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   740
  contrived: there are much simpler methods available in Isabelle for
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   741
  implementing a tactic analysing a goal according to its topmost
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   742
  connective. These simpler methods use tactic combinators, which we will
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   743
  explain in the next section.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
   744
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   745
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   746
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   747
section {* Tactic Combinators *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   748
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   749
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   750
  The purpose of tactic combinators is to build compound tactics out of
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   751
  smaller tactics. In the previous section we already used @{ML THEN}, which
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   752
  just strings together two tactics in a sequence. For example:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   755
lemma shows "(Foo \<and> Bar) \<and> False"
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   756
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
   757
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   758
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   759
       \end{minipage} *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   760
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   761
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   762
text {*
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   763
  If you want to avoid the hard-coded subgoal addressing, then, as 
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   764
  seen earlier, 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
   765
  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
   766
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   767
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   768
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
   769
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
   770
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   771
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   772
       \end{minipage} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   773
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   775
text {* 
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   776
  Here you have to specify the subgoal of interest only once and
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   777
  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
   778
  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
   779
  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
   780
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   781
  If there is a list of tactics that should all be tried out in 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   782
  sequence, you can use the combinator @{ML_ind  EVERY'}. For example
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   783
  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
   784
  be written as:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   785
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   786
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   787
ML{*val foo_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
   788
                        atac, rtac @{thm disjI1}, atac]*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   789
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   790
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   791
  There is even another way of implementing this tactic: in automatic proof
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   792
  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
   793
  are often long lists of tactics that are applied to the first
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   794
  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
   795
  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
   796
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   797
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   798
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
   799
                       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
   800
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   801
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   802
  and call @{ML foo_tac1}.  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   803
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   804
  With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML_ind  EVERY1} it must be
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   805
  guaranteed that all component tactics successfully apply; otherwise the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   806
  whole tactic will fail. If you rather want to try out a number of tactics,
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   807
  then you can use the combinator @{ML_ind  ORELSE'} for two tactics, and @{ML_ind
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   808
   FIRST'} (or @{ML_ind  FIRST1}) for a list of tactics. For example, the tactic
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   809
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   810
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   811
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   812
ML{*val orelse_xmp_tac = rtac @{thm disjI1} ORELSE' rtac @{thm conjI}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   813
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   814
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   815
  will first try out whether rule @{text disjI} applies and in case of failure 
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   816
  will try @{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
   817
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   818
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   819
lemma shows "True \<and> False" and "Foo \<or> Bar"
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   820
apply(tactic {* orelse_xmp_tac 2 *})
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   821
apply(tactic {* orelse_xmp_tac 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   822
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
   823
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   824
       \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   825
       @{subgoals [display]} 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   826
       \end{minipage}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   827
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   829
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   830
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   832
  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
   833
  as follows:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   834
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   835
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   836
ML{*val select_tac' = FIRST' [rtac @{thm conjI}, rtac @{thm impI}, 
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   837
                          rtac @{thm notI}, rtac @{thm allI}, K all_tac]*}text_raw{*\label{tac:selectprime}*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   838
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   839
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
   840
  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
   841
  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
   842
  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
   843
  @{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
   844
  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
   845
*}
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 {* select_tac' 4 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   849
apply(tactic {* select_tac' 3 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   850
apply(tactic {* select_tac' 2 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   851
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   852
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   853
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   854
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   855
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   856
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   857
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   858
  Since such repeated applications of a tactic to the reverse order of 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   859
  \emph{all} subgoals is quite common, there is the tactic combinator 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   860
  @{ML_ind  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
   861
  write: *}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   862
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   863
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
   864
apply(tactic {* ALLGOALS select_tac' *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   865
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   866
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   867
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   868
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   869
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   870
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   871
  Remember that we chose to implement @{ML select_tac'} so that it 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   872
  always  succeeds by adding @{ML all_tac} at the end of the tactic
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   873
  list. The same can be achieved with the tactic combinator @{ML_ind  TRY}.
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   874
  For example:
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   875
*}
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   876
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   877
ML{*val select_tac'' = TRY o FIRST' [rtac @{thm conjI}, rtac @{thm impI}, 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   878
                                 rtac @{thm notI}, rtac @{thm allI}]*}
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   879
text_raw{*\label{tac:selectprimeprime}*}
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   880
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   881
text {*
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   882
  This tactic behaves in the same way as @{ML select_tac'}: it tries out
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   883
  one of the given tactics and if none applies leaves the goal state 
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   884
  unchanged. This, however, can be potentially very confusing when visible to 
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   885
  the user, for example,  in cases where the goal is the form
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   886
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   887
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   888
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   889
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   890
apply(tactic {* select_tac' 1 *})
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   891
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   892
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   893
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   894
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   895
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   896
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   897
  In this case no rule applies, but because of @{ML TRY} or the inclusion of @{ML all_tac}
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   898
  the tactics do not fail. The problem with this is that for the user there is little 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   899
  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
   900
  therefore, tactics visible to the user should either change something or fail.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   901
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   902
  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
   903
  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
   904
  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
   905
  let us suppose that this deletion is \emph{not} an option. In such cases, you can
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   906
  use the combinator @{ML_ind  CHANGED} to make sure the subgoal has been changed
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   907
  by the tactic. Because now
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   908
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   909
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   910
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   911
lemma shows "E \<Longrightarrow> F"
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   912
apply(tactic {* CHANGED (select_tac' 1) *})(*<*)?(*>*)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   913
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
   914
  \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
   915
  @{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
   916
  @{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
   917
  \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
   918
*}(*<*)oops(*>*)
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   919
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   920
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   921
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   922
  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
   923
  connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   924
  completely. For this you can use the tactic combinator @{ML_ind  REPEAT}. As an example 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   925
  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
   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
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   928
ML{*val repeat_xmp_tac = REPEAT (CHANGED (select_tac' 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
   929
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   930
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
   931
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   932
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   933
apply(tactic {* repeat_xmp_tac *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   934
txt{* produces
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   935
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   936
      \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
   937
      @{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
   938
      \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
   939
(*<*)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
   940
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   941
text {*
8bea3f74889d 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
  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
   943
  because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   944
  tactic as long as it succeeds). The function
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   945
  @{ML_ind  REPEAT1} is similar, but runs the tactic at least once (failing if 
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
   946
  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
   947
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   948
  If you are after the ``primed'' version of @{ML repeat_xmp_tac}, then you 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   949
  can 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
   950
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   951
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   952
ML{*val repeat_xmp_tac' = REPEAT o CHANGED o select_tac'*}
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
8bea3f74889d 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
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
   955
  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
   956
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   957
  If you look closely at the goal state above, then you see the tactics @{ML repeat_xmp_tac}
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   958
  and @{ML repeat_xmp_tac'} are not yet quite what we are after: the problem is
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   959
  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
   960
  is applied repeatedly only to the first subgoal. To analyse also all
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   961
  resulting subgoals, you can use the tactic combinator @{ML_ind  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
   962
  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
   963
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   964
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   965
ML{*val repeat_all_new_xmp_tac = REPEAT_ALL_NEW (CHANGED o select_tac')*}
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
   966
8bea3f74889d 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
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   968
  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
   969
*}
8bea3f74889d 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
8bea3f74889d 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
lemma shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   972
apply(tactic {* repeat_all_new_xmp_tac 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
   973
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
   974
      @{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
   975
      \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
   976
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   977
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   978
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   979
  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
   980
  include in @{ML select_tac'}. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   981
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   982
  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
   983
  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
   984
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   985
*}
8bea3f74889d 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 {* etac @{thm disjE} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   989
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
   990
      
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
   991
      \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
   992
      @{subgoals [display]}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   993
      \end{minipage}\smallskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   994
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   995
      After typing
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   996
  *}
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
   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
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
   999
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
  1000
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
  1001
(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1002
back
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1003
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
  1004
8bea3f74889d 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
      \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
  1006
      @{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
  1007
      \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
  1008
(*<*)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
  1009
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1010
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
  1011
  Sometimes this leads to confusing behaviour of tactics and also has
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1012
  the potential to explode the search space for tactics.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1013
  These problems can be avoided by prefixing the tactic with the tactic 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1014
  combinator @{ML_ind  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
  1015
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1016
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1017
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
  1018
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
  1019
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
  1020
      @{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
  1021
      \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
  1022
(*<*)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
  1023
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
  1024
  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
  1025
  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
  1026
  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
  1027
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1028
  \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
  1029
  @{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
  1030
  @{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
  1031
  \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
  1032
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1033
  Recall that we implemented @{ML select_tac'} on Page~\pageref{tac:selectprime} specifically 
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1034
  so that it always succeeds. We achieved this by adding at the end the tactic @{ML all_tac}.
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1035
  We can achieve this also by using the combinator @{ML TRY}. Suppose, for example the 
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1036
  tactic
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1037
*}
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1038
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1039
ML{*val select_tac'' = FIRST' [rtac @{thm conjI}, rtac @{thm impI}, 
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1040
                           rtac @{thm notI}, rtac @{thm allI}]*}
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1041
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1042
text {*
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1043
  which will fail if none of the rules applies. However, if you prefix it as follows
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1044
*}
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1045
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1046
ML{*val select_tac''' = TRY o select_tac''*}
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1047
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1048
text {*
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1049
  then the tactic @{ML select_tac''} will be tried out and any failure is harnessed. 
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1050
  We again have to use the construction with \mbox{@{text "TRY o ..."}} since there is 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1051
  no primed version of @{ML_ind  TRY}. The tactic combinator @{ML_ind  TRYALL} will try out
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1052
  a tactic on all subgoals. For example the tactic 
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1053
*}
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1054
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1055
ML{*val triv_tac = TRYALL (rtac @{thm TrueI} ORELSE' etac @{thm FalseE})*}
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1056
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1057
text {*
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1058
  will solve all trivial subgoals involving @{term True} or @{term "False"}.
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1059
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1060
  (FIXME: say something about @{ML_ind  COND} and COND')
307
f4fa6540e280 suggestions by Peter Homeier
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
  1061
f4fa6540e280 suggestions by Peter Homeier
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
  1062
  (FIXME: PARALLEL-CHOICE PARALLEL-GOALS)
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1063
  
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
  1064
  \begin{readmore}
289
08ffafe2585d adapted to changes in Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 288
diff changeset
  1065
  Most tactic combinators described in this section are defined in @{ML_file "Pure/tactical.ML"}.
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1066
  Some combinators for the purpose of proof search are implemented in @{ML_file "Pure/search.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
  1067
  \end{readmore}
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1068
*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1069
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1070
text {*
313
1ca2f41770cc polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
  1071
  \begin{exercise}\label{ex:dyckhoff}
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1072
  Dyckhoff presents in \cite{Dyckhoff92} inference rules for a sequent
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1073
  calculus, called G4ip, in which no contraction rule is needed in order to be
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1074
  complete. As a result the rules applied in any order give a simple decision
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1075
  procedure for propositional intuitionistic logic. His rules are
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1076
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1077
  \begin{center}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1078
  \def\arraystretch{2.3}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1079
  \begin{tabular}{cc}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1080
  \infer[Ax]{A,\varGamma \Rightarrow A}{} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1081
  \infer[False]{False,\varGamma \Rightarrow G}{}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1082
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1083
  \infer[\wedge_L]{A \wedge B, \varGamma \Rightarrow G}{A, B, \varGamma \Rightarrow G} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1084
  \infer[\wedge_R]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1085
  {\varGamma \Rightarrow A\wedge B}{\varGamma \Rightarrow A & \varGamma \Rightarrow B}\\
313
1ca2f41770cc polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
  1086
  
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1087
  \infer[\vee_L]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1088
  {A\vee B, \varGamma \Rightarrow G}{A,\varGamma \Rightarrow G & B,\varGamma \Rightarrow G} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1089
  \infer[\vee_{R_1}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1090
  {\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow A} \hspace{3mm}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1091
  \infer[\vee_{R_2}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1092
  {\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow B}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1093
  
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1094
  \infer[\longrightarrow_{L_1}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1095
  {A\longrightarrow B, A, \varGamma \Rightarrow G}{B, A, \varGamma \Rightarrow G} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1096
  \infer[\longrightarrow_R]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1097
  {\varGamma \Rightarrow A\longrightarrow B}{A,\varGamma \Rightarrow B}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1098
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1099
  \infer[\longrightarrow_{L_2}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1100
  {(C \wedge D)\longrightarrow B, \varGamma \Rightarrow G}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1101
  {C\longrightarrow (D \longrightarrow B), \varGamma \Rightarrow G} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1102
  
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1103
  \infer[\longrightarrow_{L_3}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1104
  {(C \vee D)\longrightarrow B, \varGamma \Rightarrow G}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1105
  {C\longrightarrow B, D\longrightarrow B, \varGamma \Rightarrow G}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1106
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1107
  \multicolumn{2}{c}{
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1108
  \infer[\longrightarrow_{L_4}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1109
  {(C \longrightarrow D)\longrightarrow B, \varGamma \Rightarrow G}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1110
  {D\longrightarrow B, \varGamma \Rightarrow C \longrightarrow D &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1111
  B, \varGamma \Rightarrow G}}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1112
  \end{tabular}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1113
  \end{center}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1114
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1115
  Implement a tactic that explores all possibilites of applying these rules to
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1116
  a propositional formula until a goal state is reached in which all subgoals
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1117
  are discharged.  Note that in Isabelle the left-rules need to be implemented
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1118
  as elimination rules. You need to prove separate lemmas corresponding to
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1119
  $\longrightarrow_{L_{1..4}}$.  In order to explore all possibilities of applying 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1120
  the rules, use the tactic combinator @{ML_ind  DEPTH_SOLVE}, which searches 
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1121
  for a state in which all subgoals are solved. Add also rules for equality and
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1122
  run your tactic on the de Bruijn formulae discussed in Exercise~\ref{ex:debruijn}. 
313
1ca2f41770cc polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
  1123
  \end{exercise}
1ca2f41770cc polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
  1124
333
6dea46b9d2da added fixme
Christian Urban <urbanc@in.tum.de>
parents: 332
diff changeset
  1125
  \footnote{\bf FIXME: explain @{ML_ind Cong_Tac.cong_tac}}
6dea46b9d2da added fixme
Christian Urban <urbanc@in.tum.de>
parents: 332
diff changeset
  1126
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1127
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1128
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1129
section {* Simplifier Tactics *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1130
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1131
text {*
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1132
  A lot of convenience in the reasoning with Isabelle derives from its
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1133
  powerful simplifier. The power of the simplifier is a strength and a weakness at
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1134
  the same time, because you can easily make the simplifier run into a loop and
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1135
  in general its
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1136
  behaviour can be difficult to predict. There is also a multitude
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1137
  of options that you can configure to control the behaviour of the simplifier. 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1138
  We describe some of them in this and the next section.
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1139
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1140
  There are the following five main tactics behind 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1141
  the simplifier (in parentheses is their user-level counterpart):
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1142
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1143
  \begin{isabelle}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1144
  \begin{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1145
  \begin{tabular}{l@ {\hspace{2cm}}l}
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1146
   @{ML_ind  simp_tac}            & @{text "(simp (no_asm))"} \\
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1147
   @{ML_ind  asm_simp_tac}        & @{text "(simp (no_asm_simp))"} \\
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1148
   @{ML_ind  full_simp_tac}       & @{text "(simp (no_asm_use))"} \\
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1149
   @{ML_ind  asm_lr_simp_tac}     & @{text "(simp (asm_lr))"} \\
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1150
   @{ML_ind  asm_full_simp_tac}   & @{text "(simp)"}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1151
  \end{tabular}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1152
  \end{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1153
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1154
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1155
  All of the tactics take a simpset and an integer as argument (the latter as usual 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1156
  to specify  the goal to be analysed). So the proof
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1157
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1158
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1159
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
  1160
apply(simp)
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1161
done
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1162
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1163
text {*
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1164
  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
  1165
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1166
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1167
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
  1168
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
  1169
done
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1170
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1171
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1172
  If the simplifier cannot make any progress, then it leaves the goal unchanged,
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1173
  i.e., does not raise any error message. That means if you use it to unfold a 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1174
  definition for a constant and this constant is not present in the goal state, 
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1175
  you can still 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
  1176
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
  1177
  (FIXME: show rewriting of cterms)
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
  1178
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
  1179
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1180
  When using the simplifier, the crucial information you have to provide is
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1181
  the simpset. If this information is not handled with care, then the
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1182
  simplifier can easily run into a loop. Therefore a good rule of thumb is to
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1183
  use simpsets that are as minimal as possible. It might be surprising that a
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1184
  simpset is more complex than just a simple collection of theorems used as
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1185
  simplification rules. One reason for the complexity is that the simplifier
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1186
  must be able to rewrite inside terms and should also be able to rewrite
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1187
  according to rules that have preconditions.
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1188
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1189
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1190
  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
  1191
  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
  1192
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1193
  \begin{isabelle}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1194
  \begin{center}
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1195
  \mbox{\inferrule{@{text "t\<^isub>1 \<equiv> s\<^isub>1 \<dots> t\<^isub>n \<equiv> s\<^isub>n"}}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1196
                  {@{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
  1197
  \end{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1198
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1199
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1200
  with @{text "constr"} being a constant, like @{const "If"} or @{const "Let"}. 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1201
  Every simpset contains only
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1202
  one congruence rule for each term-constructor, which however can be
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1203
  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
  1204
  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
  1205
  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
  1206
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
  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
  1209
  Isabelle called a subgoaler, a solver and a looper. These can be arbitrary
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1210
  tactics that can be installed in a simpset and which are called at
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1211
  various stages during simplification. However, simpsets also include
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1212
  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
  1213
  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
  1214
  the ``then'' and ``else'' branches of if-statements under the corresponding
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1215
  preconditions.
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1216
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1217
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1218
  \begin{readmore}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1219
  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
  1220
  and @{ML_file "Pure/simplifier.ML"}. The simplifier for HOL is set up in 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1221
  @{ML_file "HOL/Tools/simpdata.ML"}. The generic splitter is implemented in 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1222
  @{ML_file  "Provers/splitter.ML"}.
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1223
  \end{readmore}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1224
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
  1225
  \begin{readmore}
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1226
  FIXME: Find the right place: Discrimination nets are implemented
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
  1227
  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
  1228
  \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
  1229
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1230
  The most common combinators to modify simpsets are:
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1231
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1232
  \begin{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1233
  \begin{tabular}{ll}
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1234
  @{ML_ind  addsimps}    & @{ML_ind  delsimps}\\
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1235
  @{ML_ind  addcongs}    & @{ML_ind  delcongs}\\
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1236
  @{ML_ind  addsimprocs} & @{ML_ind  delsimprocs}\\
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1237
  \end{tabular}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1238
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1239
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1240
  (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
  1241
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1242
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1243
text_raw {*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1244
\begin{figure}[t]
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1245
\begin{minipage}{\textwidth}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1246
\begin{isabelle}*}
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1247
ML{*fun print_ss ctxt ss =
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1248
let
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1249
  val {simps, congs, procs, ...} = Simplifier.dest_ss ss
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1250
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1251
  fun name_thm (nm, thm) =
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 243
diff changeset
  1252
      "  " ^ nm ^ ": " ^ (string_of_thm_no_vars ctxt thm)
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1253
  fun name_ctrm (nm, ctrm) =
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 243
diff changeset
  1254
      "  " ^ nm ^ ": " ^ (string_of_cterms ctxt ctrm)
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1255
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1256
  val s = ["Simplification rules:"] @ map name_thm simps @
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1257
          ["Congruences rules:"] @ map name_thm congs @
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1258
          ["Simproc patterns:"] @ map name_ctrm procs
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1259
in
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1260
  s |> cat_lines
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
  1261
    |> tracing
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1262
end*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1263
text_raw {* 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1264
\end{isabelle}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1265
\end{minipage}
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1266
\caption{The function @{ML_ind  dest_ss in Simplifier} returns a record containing
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1267
  all printable information stored in a simpset. We are here only interested in the 
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1268
  simplification rules, congruence rules and simprocs.\label{fig:printss}}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1269
\end{figure} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1270
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
  1271
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1272
text {* 
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1273
  To see how they work, consider the function in Figure~\ref{fig:printss}, which
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1274
  prints out some parts of a simpset. If you use it to print out the components of the
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1275
  empty simpset, i.e., @{ML_ind  empty_ss}
157
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
  @{ML_response_fake [display,gray]
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1278
  "print_ss @{context} empty_ss"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1279
"Simplification rules:
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1280
Congruences rules:
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1281
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1282
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1283
  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
  1284
  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
  1285
  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
  1286
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1287
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1288
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
  1289
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1290
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1291
  Printing then out the components of the simpset gives:
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1292
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1293
  @{ML_response_fake [display,gray]
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1294
  "print_ss @{context} ss1"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1295
"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
  1296
  ??.unknown: A - B \<inter> C \<equiv> A - B \<union> (A - C)
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1297
Congruences rules:
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1298
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1299
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1300
  (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
  1301
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1302
  Adding also 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
  1303
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1304
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1305
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
  1306
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1307
text {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1308
  gives
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1309
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1310
  @{ML_response_fake [display,gray]
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1311
  "print_ss @{context} ss2"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1312
"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
  1313
  ??.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
  1314
Congruences rules:
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1315
  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
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1316
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1317
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1318
  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
  1319
  expects this form of the simplification and congruence rules. However, even 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1320
  when adding these lemmas to @{ML empty_ss} we do not end up with anything useful yet.
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1321
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1322
  In the context of HOL, the first really useful simpset is @{ML_ind  HOL_basic_ss}. While
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1323
  printing out the components of this simpset
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1324
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1325
  @{ML_response_fake [display,gray]
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1326
  "print_ss @{context} HOL_basic_ss"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1327
"Simplification rules:
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1328
Congruences rules:
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1329
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1330
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1331
  also produces ``nothing'', the printout is misleading. In fact
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1332
  the @{ML HOL_basic_ss} is setup so that it can solve goals of the
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1333
  form  
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1334
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1335
  \begin{isabelle}
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1336
  @{thm TrueI}, @{thm refl[no_vars]}, @{term "t \<equiv> t"} and @{thm FalseE[no_vars]}; 
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1337
  \end{isabelle}
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1338
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1339
  and also resolve with assumptions. For example:
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1340
*}
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 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1343
 "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
  1344
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
  1345
done
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 {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1348
  This behaviour is not because of simplification rules, but how the subgoaler, solver 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1349
  and looper are set up in @{ML_ind  HOL_basic_ss}.
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1350
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1351
  The simpset @{ML_ind  HOL_ss} is an extension of @{ML HOL_basic_ss} containing 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1352
  already many useful simplification and congruence rules for the logical 
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1353
  connectives in HOL. 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1354
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1355
  @{ML_response_fake [display,gray]
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1356
  "print_ss @{context} HOL_ss"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1357
"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
  1358
  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
  1359
  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
  1360
  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
  1361
  \<dots>
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1362
Congruences rules:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1363
  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
  1364
    \<Longrightarrow> (PROP P =simp=> PROP Q) \<equiv> (PROP P' =simp=> PROP Q')
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1365
  op -->: \<lbrakk>P \<equiv> P'; P' \<Longrightarrow> Q \<equiv> Q'\<rbrakk> \<Longrightarrow> P \<longrightarrow> Q \<equiv> P' \<longrightarrow> Q'
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1366
Simproc patterns:
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
  1367
  \<dots>"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1368
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1369
  
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1370
  The simplifier is often used to unfold definitions in a proof. For this the
331
46100dc4a808 used rewrite_goal_tac (instead of rewrite_goals_tac)
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1371
  simplifier implements the tactic @{ML_ind  rewrite_goal_tac}.\footnote{\bf FIXME: 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1372
  see LocalDefs infrastructure.} Suppose for example the
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1373
  definition
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1374
*}
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1375
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1376
definition "MyTrue \<equiv> True"
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1377
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1378
text {*
331
46100dc4a808 used rewrite_goal_tac (instead of rewrite_goals_tac)
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1379
  then we can use this tactic to unfold the definition of the constant.
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1380
*}
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1381
331
46100dc4a808 used rewrite_goal_tac (instead of rewrite_goals_tac)
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1382
lemma shows "MyTrue \<Longrightarrow> True"
46100dc4a808 used rewrite_goal_tac (instead of rewrite_goals_tac)
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1383
apply(tactic {* rewrite_goal_tac @{thms MyTrue_def} 1 *})
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1384
txt{* producing the goal state
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1385
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1386
      \begin{minipage}{\textwidth}
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1387
      @{subgoals [display]}
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1388
      \end{minipage} *}
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1389
(*<*)oops(*>*)
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1390
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1391
text {*
331
46100dc4a808 used rewrite_goal_tac (instead of rewrite_goals_tac)
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1392
  If you want to unfold definitions in all subgoals, then use the 
46100dc4a808 used rewrite_goal_tac (instead of rewrite_goals_tac)
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1393
  the tactic @{ML_ind rewrite_goals_tac}.
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1394
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1395
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1396
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1397
text_raw {*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1398
\begin{figure}[p]
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1399
\begin{boxedminipage}{\textwidth}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1400
\begin{isabelle} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1401
types  prm = "(nat \<times> nat) list"
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1402
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
  1403
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1404
overloading 
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1405
  perm_nat \<equiv> "perm :: prm \<Rightarrow> nat \<Rightarrow> nat"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1406
  perm_prod \<equiv> "perm :: prm \<Rightarrow> ('a\<times>'b) \<Rightarrow> ('a\<times>'b)"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1407
  perm_list \<equiv> "perm :: prm \<Rightarrow> 'a list \<Rightarrow> 'a list"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1408
begin
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1409
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1410
fun swap::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1411
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1412
 "swap a b c = (if c=a then b else (if c=b then a else c))"
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1413
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1414
primrec perm_nat
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1415
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1416
  "perm_nat [] c = c"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1417
| "perm_nat (ab#pi) c = swap (fst ab) (snd ab) (perm_nat pi c)"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1418
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1419
fun perm_prod 
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1420
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1421
 "perm_prod pi (x, y) = (pi\<bullet>x, pi\<bullet>y)"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1422
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1423
primrec perm_list
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1424
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1425
  "perm_list pi [] = []"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1426
| "perm_list pi (x#xs) = (pi\<bullet>x)#(perm_list pi xs)"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1427
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1428
end
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1429
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1430
lemma perm_append[simp]:
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1431
fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1432
shows "((pi\<^isub>1@pi\<^isub>2)\<bullet>c) = (pi\<^isub>1\<bullet>(pi\<^isub>2\<bullet>c))"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1433
by (induct pi\<^isub>1) (auto) 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1434
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1435
lemma perm_bij[simp]:
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1436
fixes c d::"nat" and pi::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1437
shows "(pi\<bullet>c = pi\<bullet>d) = (c = d)" 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1438
by (induct pi) (auto)
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1439
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1440
lemma perm_rev[simp]:
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1441
fixes c::"nat" and pi::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1442
shows "pi\<bullet>((rev pi)\<bullet>c) = c"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1443
by (induct pi arbitrary: c) (auto)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1444
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1445
lemma perm_compose:
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1446
fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1447
shows "pi\<^isub>1\<bullet>(pi\<^isub>2\<bullet>c) = (pi\<^isub>1\<bullet>pi\<^isub>2)\<bullet>(pi\<^isub>1\<bullet>c)" 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1448
by (induct pi\<^isub>2) (auto)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1449
text_raw {*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1450
\end{isabelle}
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1451
\end{boxedminipage}
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1452
\caption{A simple theory about permutations over @{typ nat}s. The point is that the
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1453
  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
  1454
  it would cause the simplifier to loop. It can still be used as a simplification 
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1455
  rule if the permutation in the right-hand side is sufficiently protected.
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1456
  \label{fig:perms}}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1457
\end{figure} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1458
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1459
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1460
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1461
  The simplifier is often used in order to bring terms into a normal form.
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1462
  Unfortunately, often the situation arises that the corresponding
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1463
  simplification rules will cause the simplifier to run into an infinite
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1464
  loop. Consider for example the simple theory about permutations over natural
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1465
  numbers shown in Figure~\ref{fig:perms}. The purpose of the lemmas is to
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1466
  push permutations as far inside as possible, where they might disappear by
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1467
  Lemma @{thm [source] perm_rev}. However, to fully normalise all instances,
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1468
  it would be desirable to add also the lemma @{thm [source] perm_compose} to
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1469
  the simplifier for pushing permutations over other permutations. Unfortunately, 
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1470
  the right-hand side of this lemma is again an instance of the left-hand side 
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1471
  and so causes an infinite loop. There seems to be no easy way to reformulate 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1472
  this rule and so one ends up with clunky proofs like:
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1473
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1474
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1475
lemma 
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1476
fixes c d::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1477
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)"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1478
apply(simp)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1479
apply(rule trans)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1480
apply(rule perm_compose)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1481
apply(simp)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1482
done 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1483
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1484
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1485
  It is however possible to create a single simplifier tactic that solves
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1486
  such proofs. The trick is to introduce an auxiliary constant for permutations 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1487
  and split the simplification into two phases (below actually three). Let 
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1488
  assume the auxiliary constant is
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1489
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1490
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1491
definition
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1492
  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
  1493
where
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1494
  "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
  1495
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1496
text {* Now the two lemmas *}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1497
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1498
lemma perm_aux_expand:
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1499
fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1500
shows "pi\<^isub>1\<bullet>(pi\<^isub>2\<bullet>c) = pi\<^isub>1 \<bullet>aux (pi\<^isub>2\<bullet>c)" 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1501
unfolding perm_aux_def by (rule refl)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1502
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1503
lemma perm_compose_aux:
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1504
fixes c::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1505
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)" 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1506
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
  1507
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1508
text {* 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1509
  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
  1510
  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
  1511
  added to the simplifier, because now the right-hand side is not anymore an instance 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1512
  of the left-hand side. In a sense it freezes all redexes of permutation compositions
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1513
  after one step. In this way, we can split simplification of permutations
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
  1514
  into three phases without the user noticing anything about the auxiliary 
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1515
  constant. We first freeze any instance of permutation compositions in the term using 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1516
  lemma @{thm [source] "perm_aux_expand"} (Line 9);
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1517
  then simplify all other permutations including pushing permutations over
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1518
  other permutations by rule @{thm [source] perm_compose_aux} (Line 10); and
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1519
  finally ``unfreeze'' all instances of permutation compositions by unfolding 
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1520
  the definition of the auxiliary constant. 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1521
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1522
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1523
ML %linenosgray{*val perm_simp_tac =
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1524
let
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1525
  val thms1 = [@{thm perm_aux_expand}]
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1526
  val thms2 = [@{thm perm_append}, @{thm perm_bij}, @{thm perm_rev}, 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1527
               @{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
  1528
               @{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
  1529
  val thms3 = [@{thm perm_aux_def}]
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1530
in
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1531
  simp_tac (HOL_basic_ss addsimps thms1)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1532
  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
  1533
  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
  1534
end*}
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1535
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1536
text {*
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1537
  For all three phases we have to build simpsets adding specific lemmas. As is sufficient
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1538
  for our purposes here, we can add these lemmas to @{ML HOL_basic_ss} in order to obtain
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1539
  the desired results. Now we can solve the following lemma
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1540
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1541
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1542
lemma 
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1543
fixes c d::"nat" and pi\<^isub>1 pi\<^isub>2::"prm"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1544
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)"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1545
apply(tactic {* perm_simp_tac 1 *})
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1546
done
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1547
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1548
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1549
text {*
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1550
  in one step. This tactic can deal with most instances of normalising permutations.
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1551
  In order to solve all cases we have to deal with corner-cases such as the
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1552
  lemma being an exact instance of the permutation composition lemma. This can
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1553
  often be done easier by implementing a simproc or a conversion. Both will be 
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1554
  explained in the next two chapters.
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1555
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1556
  (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
  1557
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1558
  (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
  1559
  do with weak congruence constants?)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1560
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1561
  (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
  1562
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1563
  (FIXME: @{ML ObjectLogic.full_atomize_tac}, 
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1564
  @{ML ObjectLogic.rulify_tac})
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1565
240
d111f5988e49 replaced explode by Symbol.explode
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
  1566
  (FIXME: what are @{ML mksimps_pairs}; used in Nominal.thy)
d111f5988e49 replaced explode by Symbol.explode
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
  1567
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 243
diff changeset
  1568
  (FIXME: explain @{ML simplify} and @{ML "Simplifier.rewrite_rule"} etc.)
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 243
diff changeset
  1569
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1570
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1571
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1572
section {* Simprocs *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1573
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1574
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1575
  In Isabelle you can also implement custom simplification procedures, called
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1576
  \emph{simprocs}. Simprocs can be triggered by the simplifier on a specified
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1577
  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
  1578
  cases where a rewriting rule must be produced on ``demand'' or when
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1579
  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
  1580
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1581
  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
  1582
  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
  1583
  you can use the function:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1584
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1585
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1586
ML %linenosgray{*fun fail_simproc simpset redex = 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1587
let
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1588
  val ctxt = Simplifier.the_context simpset
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
  1589
  val _ = tracing ("The redex: " ^ (string_of_cterm ctxt redex))
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1590
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1591
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1592
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1593
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1594
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1595
  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
  1596
  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
  1597
  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
  1598
  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
  1599
  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
  1600
  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
  1601
  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
  1602
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1603
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1604
simproc_setup %gray fail ("Suc n") = {* K fail_simproc *}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1605
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1606
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1607
  where the second argument specifies the pattern and the right-hand side
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1608
  contains the code of the simproc (we have to use @{ML K} since we are ignoring
230
8def50824320 added material about OuterKeyword.keyword and OuterParse.reserved
Christian Urban <urbanc@in.tum.de>
parents: 229
diff changeset
  1609
  an argument about morphisms. 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1610
  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
  1611
  it fires on the lemma:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1612
*}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1613
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1614
lemma shows "Suc 0 = 1"
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1615
apply(simp)
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1616
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1617
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1618
text {*
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
  1619
  \begin{isabelle}
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
  1620
  @{text "> The redex: Suc 0"}\\
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
  1621
  @{text "> The redex: Suc 0"}\\
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
  1622
  \end{isabelle}
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
  1623
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1624
  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
  1625
  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
  1626
  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
  1627
  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
  1628
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1629
  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
  1630
  \isacommand{declare}-statement. For example the simproc will be deleted
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1631
  with the declaration
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1632
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1633
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1634
declare [[simproc del: fail]]
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1635
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1636
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1637
  If you want to see what happens with just \emph{this} simproc, without any 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1638
  interference from other rewrite rules, you can call @{text fail} 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1639
  as follows:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1640
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1641
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1642
lemma shows "Suc 0 = 1"
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1643
apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [@{simproc fail}]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1644
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1645
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1646
text {*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1647
  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
  1648
  left unchanged. 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1649
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1650
  Setting up a simproc using the command \isacommand{simproc\_setup} will
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1651
  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
  1652
  want this, then you have to use a slightly different method for setting 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1653
  up the simproc. First the function @{ML fail_simproc} needs to be modified
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1654
  to
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1655
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1656
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1657
ML{*fun fail_simproc' simpset redex = 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1658
let
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1659
  val ctxt = Simplifier.the_context simpset
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1660
  val _ = tracing ("The redex: " ^ (string_of_term ctxt redex))
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1661
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1662
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1663
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1664
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1665
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1666
  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
  1667
  (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
  1668
  We can turn this function into a proper simproc using the function 
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1669
  @{ML Simplifier.simproc_i}:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1670
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1671
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1672
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1673
ML{*val fail' = 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1674
let 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1675
  val thy = @{theory}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1676
  val pat = [@{term "Suc n"}]
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1677
in
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1678
  Simplifier.simproc_i thy "fail_simproc'" pat (K fail_simproc')
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1679
end*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1680
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1681
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1682
  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
  1683
  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
  1684
  Now the simproc is set up and can be explicitly added using
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1685
  @{ML_ind  addsimprocs} to a simpset whenever
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1686
  needed. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1687
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1688
  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
  1689
  see this in the proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1690
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1691
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1692
lemma shows "Suc (Suc 0) = (Suc 1)"
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1693
apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [fail']) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1694
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1695
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1696
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1697
  The simproc @{ML fail'} prints out the sequence 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1698
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1699
@{text [display]
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1700
"> Suc 0
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1701
> Suc (Suc 0) 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1702
> Suc 1"}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1703
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1704
  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
  1705
  rewrites terms according to the equation:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1706
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1707
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1708
lemma plus_one: 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1709
  shows "Suc n \<equiv> n + 1" by simp
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1710
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1711
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1712
  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
  1713
  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
  1714
  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
  1715
  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
  1716
  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
  1717
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1718
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1719
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1720
ML{*fun plus_one_simproc ss redex =
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1721
  case redex of
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1722
    @{term "Suc 0"} => NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1723
  | _ => SOME @{thm plus_one}*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1724
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1725
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1726
  and set up the simproc as follows.
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1727
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1728
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1729
ML{*val plus_one =
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1730
let
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1731
  val thy = @{theory}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1732
  val pat = [@{term "Suc n"}] 
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1733
in
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1734
  Simplifier.simproc_i thy "sproc +1" pat (K plus_one_simproc)
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1735
end*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1736
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1737
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1738
  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
  1739
  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
  1740
  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
  1741
  in the following proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1742
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1743
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1744
lemma shows "P (Suc (Suc (Suc 0))) (Suc 0)"
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1745
apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [plus_one]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1746
txt{*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1747
  where the simproc produces the goal state
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1748
  
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1749
  \begin{minipage}{\textwidth}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1750
  @{subgoals[display]}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1751
  \end{minipage}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1752
*}  
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1753
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1754
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1755
text {*
133
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
  1756
  As usual with rewriting you have to worry about looping: you already have
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1757
  a loop with @{ML plus_one}, if you apply it with the default simpset (because
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1758
  the default simpset contains a rule which just does the opposite of @{ML plus_one},
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1759
  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
  1760
  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
  1761
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1762
  Next let us implement a simproc that replaces terms of the form @{term "Suc n"}
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1763
  with the number @{text n} increased by one. First we implement a function that
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1764
  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
  1765
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1766
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1767
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
  1768
  | 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
  1769
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1770
text {* 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1771
  It uses the library function @{ML_ind  dest_number in HOLogic} that transforms
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1772
  (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
  1773
  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
  1774
  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
  1775
  @{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
  1776
*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1777
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1778
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
  1779
let
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1780
  val num = HOLogic.mk_number @{typ "nat"} n
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1781
  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
  1782
in
214
7e04dc2368b0 updated to latest Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 213
diff changeset
  1783
  Goal.prove ctxt [] [] goal (K (Arith_Data.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
  1784
end*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1785
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1786
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1787
  From the integer value it generates the corresponding number term, called 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1788
  @{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
  1789
  (Line 4), which it proves by the arithmetic tactic in Line 6. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1790
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 217
diff changeset
  1791
  For our purpose at the moment, proving the meta-equation using @{ML
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 217
diff changeset
  1792
  arith_tac in Arith_Data} is fine, but there is also an alternative employing
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 217
diff changeset
  1793
  the simplifier with a special simpset. For the kind of lemmas we
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 217
diff changeset
  1794
  want to prove here, the simpset @{text "num_ss"} should suffice.
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1795
*}
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1796
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1797
ML{*fun get_thm_alt ctxt (t, n) =
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1798
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1799
  val num = HOLogic.mk_number @{typ "nat"} n
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1800
  val goal = Logic.mk_equals (t, num)
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1801
  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
  1802
          @{thms nat_number} @ @{thms neg_simps} @ @{thms plus_nat.simps}
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1803
in
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1804
  Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1))
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1805
end*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1806
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1807
text {*
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1808
  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
  1809
  something to go wrong; in contrast it is much more difficult to predict 
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 217
diff changeset
  1810
  what happens with @{ML arith_tac in Arith_Data}, especially in more complicated 
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1811
  circumstances. The disadvantage of @{ML get_thm_alt} is to find a simpset
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1812
  that is sufficiently powerful to solve every instance of the lemmas
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1813
  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
  1814
  ``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
  1815
  way than tracing the simplifier to obtain the lemmas that are successfully 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1816
  applied during simplification. Alas, there is none.} 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1817
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1818
  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
  1819
  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
  1820
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1821
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1822
ML{*fun nat_number_simproc ss t =
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1823
let 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1824
  val ctxt = Simplifier.the_context ss
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1825
in
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1826
  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
  1827
  handle TERM _ => NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1828
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1829
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1830
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1831
  This function uses the fact that @{ML dest_suc_trm} might raise an exception 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1832
  @{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
  1833
  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
  1834
  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
  1835
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1836
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1837
ML{*val nat_number =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1838
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1839
  val thy = @{theory}
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1840
  val pat = [@{term "Suc n"}]
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1841
in 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1842
  Simplifier.simproc_i thy "nat_number" pat (K nat_number_simproc) 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1843
end*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1844
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1845
text {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1846
  Now in the lemma
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1847
  *}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1848
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1849
lemma "P (Suc (Suc 2)) (Suc 99) (0::nat) (Suc 4 + Suc 0) (Suc (0 + 0))"
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1850
apply(tactic {* simp_tac (HOL_ss addsimprocs [nat_number]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1851
txt {* 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1852
  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
  1853
  
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1854
  \begin{minipage}{\textwidth}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1855
  @{subgoals [display]}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1856
  \end{minipage}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1857
  *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1858
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1859
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1860
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1861
  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
  1862
  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
  1863
  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
  1864
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1865
  \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
  1866
  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
  1867
  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
  1868
  @{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
  1869
  \end{exercise}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1870
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1871
  (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
  1872
  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
  1873
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1874
137
a9685909944d new pfd file
Christian Urban <urbanc@in.tum.de>
parents: 135
diff changeset
  1875
section {* Conversions\label{sec:conversion} *}
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1876
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1877
text {*
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1878
  Conversions are a thin layer on top of Isabelle's inference kernel, and 
169
d3fcc1a0272c Corrected small mistake.
boehmes
parents: 149
diff changeset
  1879
  can be viewed 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
  1880
  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
  1881
  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
  1882
  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
  1883
  to theorems via tactics. The type for conversions is
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1884
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1885
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1886
ML{*type conv = cterm -> thm*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1887
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1888
text {*
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1889
  whereby the produced theorem is always a meta-equality. A simple conversion
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1890
  is the function @{ML_ind  all_conv in Conv}, which maps a @{ML_type cterm} to an
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1891
  instance of the (meta)reflexivity theorem. For example:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1892
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1893
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1894
  "Conv.all_conv @{cterm \"Foo \<or> Bar\"}"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1895
  "Foo \<or> Bar \<equiv> Foo \<or> Bar"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1896
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1897
  Another simple conversion is @{ML_ind  no_conv in Conv} which always raises the 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1898
  exception @{ML CTERM}.
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1899
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1900
  @{ML_response_fake [display,gray]
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1901
  "Conv.no_conv @{cterm True}" 
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1902
  "*** Exception- CTERM (\"no conversion\", []) raised"}
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1903
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1904
  A more interesting conversion is the function @{ML_ind  beta_conversion in Thm}: 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
  1905
  produces a meta-equation between a term and its beta-normal form. For example
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1906
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1907
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1908
  "let
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1909
  val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1910
  val two = @{cterm \"2::nat\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1911
  val ten = @{cterm \"10::nat\"}
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1912
  val ctrm = Thm.capply (Thm.capply add two) ten
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1913
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1914
  Thm.beta_conversion true ctrm
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1915
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1916
  "((\<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
  1917
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1918
  If you run this example, you will notice that the actual response is the 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1919
  seemingly nonsensical @{term
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1920
  "2 + 10 \<equiv> 2 + (10::nat)"}. The reason is that the pretty-printer for
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1921
  @{ML_type cterm}s eta-normalises terms and therefore produces this output.
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1922
  If we get hold of the ``raw'' representation of the produced theorem, 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1923
  we obtain the expected result.
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1924
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1925
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1926
  @{ML_response [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1927
"let
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1928
  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
  1929
  val two = @{cterm \"2::nat\"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1930
  val ten = @{cterm \"10::nat\"}
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1931
  val ctrm = 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
  1932
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1933
  Thm.prop_of (Thm.beta_conversion true ctrm)
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1934
end"
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1935
"Const (\"==\",\<dots>) $ 
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1936
  (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
  1937
    (Const (\"HOL.plus_class.plus\",\<dots>) $ \<dots> $ \<dots>)"} 
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  1938
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1939
  The argument @{ML true} in @{ML beta_conversion in Thm} indicates that 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1940
  the right-hand side should be fully beta-normalised. If instead 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1941
  @{ML false} is given, then only a single beta-reduction is performed 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1942
  on the outer-most level. 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1943
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1944
  The main point of conversions is that they can be used for rewriting
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1945
  @{ML_type cterm}s. One example is the function 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1946
  @{ML_ind  rewr_conv in Conv}, which expects a meta-equation as an 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1947
  argument. Suppose the following meta-equation.
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1948
  
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1949
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1950
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1951
lemma true_conj1: "True \<and> P \<equiv> P" by simp
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  1952
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1953
text {* 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1954
  It can be used for example to rewrite @{term "True \<and> (Foo \<longrightarrow> Bar)"} 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1955
  to @{term "Foo \<longrightarrow> Bar"}. The code is as follows.
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1956
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1957
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1958
"let 
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1959
  val ctrm = @{cterm \"True \<and> (Foo \<longrightarrow> Bar)\"}
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1960
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1961
  Conv.rewr_conv @{thm true_conj1} ctrm
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1962
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1963
  "True \<and> (Foo \<longrightarrow> Bar) \<equiv> Foo \<longrightarrow> Bar"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1964
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1965
  Note, however, that the function @{ML_ind  rewr_conv in 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
  1966
  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
  1967
  exactly the 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1968
  left-hand side of the theorem, then  @{ML_ind  rewr_conv in Conv} fails, raising 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1969
  the exception @{ML CTERM}.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1970
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1971
  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
  1972
  combining several conversions into one. For this you can use conversion
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1973
  combinators. The simplest conversion combinator is @{ML_ind  then_conv}, 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1974
  which applies one conversion after another. For example
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1975
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1976
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1977
"let
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1978
  val conv1 = Thm.beta_conversion false
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1979
  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
  1980
  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
  1981
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1982
  (conv1 then_conv conv2) ctrm
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1983
end"
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1984
  "(\<lambda>x. x \<and> False) True \<equiv> False"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  1985
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1986
  where we first beta-reduce the term and then rewrite according to
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1987
  @{thm [source] true_conj1}. (When running this example recall the 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  1988
  problem with the pretty-printer normalising all terms.)
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1989
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1990
  The conversion combinator @{ML_ind  else_conv} tries out the 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1991
  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
  1992
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  1993
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1994
"let
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  1995
  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
  1996
  val ctrm1 = @{cterm \"True \<and> Q\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1997
  val ctrm2 = @{cterm \"P \<or> (True \<and> Q)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1998
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  1999
  (conv ctrm1, conv ctrm2)
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2000
end"
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2001
"(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
  2002
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2003
  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
  2004
  in the first case, but fails in the second. The whole conversion
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
  2005
  does not fail, however, because the combinator @{ML else_conv in Conv} will then 
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
  2006
  try out @{ML all_conv in Conv}, which always succeeds.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2007
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2008
  The conversion combinator @{ML_ind  try_conv in Conv} constructs a conversion 
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2009
  which is tried out on a term, but in case of failure just does nothing.
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2010
  For example
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2011
  
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2012
  @{ML_response_fake [display,gray]
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2013
"let
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2014
  val conv = Conv.try_conv (Conv.rewr_conv @{thm true_conj1})
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2015
  val ctrm = @{cterm \"True \<or> P\"}
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2016
in
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2017
  conv ctrm
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2018
end"
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2019
  "True \<or> P \<equiv> True \<or> P"}
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2020
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  2021
  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
  2022
  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
  2023
  only apply to the outer-most level of a @{ML_type cterm}. In what follows we
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2024
  will lift this restriction. The combinators @{ML_ind  fun_conv in Conv} 
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2025
  and @{ML_ind  arg_conv in Conv} will apply
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2026
  a conversion to the first, respectively second, argument of an application.  
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2027
  For example
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2028
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2029
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2030
"let 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2031
  val conv = Conv.arg_conv (Conv.rewr_conv @{thm true_conj1})
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2032
  val ctrm = @{cterm \"P \<or> (True \<and> Q)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2033
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2034
  conv ctrm
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2035
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2036
"P \<or> (True \<and> Q) \<equiv> P \<or> Q"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2037
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2038
  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
  2039
  arguments.  Therefore the term must be of the form @{text "(Const \<dots> $ t1) $ t2"}. The
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2040
  conversion is then applied to @{text "t2"}, which in the example above
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2041
  stands for @{term "True \<and> Q"}. The function @{ML fun_conv in Conv} would apply
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2042
  the conversion to the term @{text "(Const \<dots> $ t1)"}.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2043
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2044
  The function @{ML_ind  abs_conv in Conv} applies a conversion under an 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2045
  abstraction. For example:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2046
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2047
  @{ML_response_fake [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2048
"let 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2049
  val conv = Conv.rewr_conv @{thm true_conj1} 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2050
  val ctrm = @{cterm \"\<lambda>P. True \<and> (P \<and> Foo)\"}
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2051
in
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2052
  Conv.abs_conv (K conv) @{context} ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2053
end"
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2054
  "\<lambda>P. True \<and> (P \<and> Foo) \<equiv> \<lambda>P. P \<and> Foo"}
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2055
  
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2056
  Note that this conversion needs a context as an argument. We also give the
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2057
  conversion as @{text "(K conv)"}, which is a function that ignores its
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2058
  argument (the argument being a sufficiently freshened version of the
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2059
  variable that is abstracted and a context). The conversion that goes under
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2060
  an application is @{ML_ind  combination_conv in Conv}. It expects two
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2061
  conversions as arguments, each of which is applied to the corresponding
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 291
diff changeset
  2062
  ``branch'' of the application. An abbreviation for this conversion is the
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2063
  function @{ML_ind  comb_conv in Conv}, which applies the same conversion
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 291
diff changeset
  2064
  to both branches.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2065
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
  2066
  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
  2067
  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
  2068
  in all possible positions.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2069
*}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2070
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2071
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
  2072
  case (Thm.term_of ctrm) of
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2073
    @{term "op \<and>"} $ @{term True} $ _ => 
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2074
      (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
  2075
         Conv.rewr_conv @{thm true_conj1}) ctrm
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 291
diff changeset
  2076
  | _ $ _ => Conv.comb_conv (all_true1_conv ctxt) ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2077
  | 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
  2078
  | _ => Conv.all_conv ctrm*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2079
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2080
text {*
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2081
  This function ``fires'' if the terms is of the form @{text "(True \<and> \<dots>)"}. 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2082
  It descends under applications (Line 6 and 7) and abstractions 
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
  2083
  (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
  2084
  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
  2085
  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
  2086
  conversion in action, consider the following example:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2087
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2088
@{ML_response_fake [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2089
"let
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2090
  val conv = all_true1_conv @{context}
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2091
  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
  2092
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2093
  conv ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2094
end"
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2095
  "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
  2096
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  2097
  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
  2098
  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
  2099
  @{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
  2100
  the conversion should be as follows.
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2101
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2102
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2103
ML{*fun if_true1_conv ctxt ctrm =
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2104
  case Thm.term_of ctrm of
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2105
    Const (@{const_name If}, _) $ _ =>
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2106
      Conv.arg_conv (all_true1_conv ctxt) ctrm
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 291
diff changeset
  2107
  | _ $ _ => Conv.comb_conv (if_true1_conv ctxt) ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2108
  | 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
  2109
  | _ => Conv.all_conv ctrm *}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2110
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2111
text {*
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  2112
  Here is an example for this conversion:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2113
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2114
  @{ML_response_fake [display,gray]
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2115
"let
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2116
  val conv = if_true1_conv @{context}
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2117
  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
  2118
       @{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
  2119
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2120
  conv ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2121
end"
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2122
"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
  2123
\<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
  2124
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2125
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2126
text {*
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2127
  So far we only applied conversions to @{ML_type cterm}s. Conversions can, however, 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2128
  also work on theorems using the function @{ML_ind  fconv_rule in Conv}. As an example, 
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  2129
  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
  2130
*}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2131
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2132
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
  2133
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2134
text {*
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2135
  Using the conversion @{ML all_true1_conv} you can transform this theorem into a 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2136
  new theorem as follows
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2137
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2138
  @{ML_response_fake [display,gray]
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2139
"let
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2140
  val conv = Conv.fconv_rule (all_true1_conv @{context}) 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2141
  val thm = @{thm foo_test}
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2142
in
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2143
  conv thm
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2144
end" 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2145
  "?P \<or> \<not> ?P"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2146
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2147
  Finally, conversions can also be turned into tactics and then applied to
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2148
  goal states. This can be done with the help of the function @{ML_ind  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
  2149
  and also some predefined conversion combinators that traverse a goal
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2150
  state. The combinators for the goal state are: 
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2151
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2152
  \begin{itemize}
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2153
  \item @{ML_ind  params_conv in Conv} for converting under parameters
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2154
  (i.e.~where goals are of the form @{text "\<And>x. P x \<Longrightarrow> Q x"})
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2155
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2156
  \item @{ML_ind  prems_conv in Conv} for applying a conversion to all
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2157
  premises of a goal, and
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2158
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2159
  \item @{ML_ind  concl_conv in Conv} for applying a conversion to the
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2160
  conclusion of a goal.
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2161
  \end{itemize}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2162
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2163
  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
  2164
  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
  2165
  Here is a tactic doing exactly that:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2166
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2167
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2168
ML{*fun true1_tac ctxt =
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  2169
  CONVERSION 
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  2170
    (Conv.params_conv ~1 (fn ctxt =>
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  2171
       (Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2172
          Conv.concl_conv ~1 (all_true1_conv ctxt))) ctxt)*}
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2173
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2174
text {* 
148
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  2175
  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
  2176
  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
  2177
  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
  2178
  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
  2179
  conclusions). To test the tactic, consider the proof
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2180
*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2181
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2182
lemma
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2183
  "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
  2184
  (if True \<and> Q then True \<and> Q else P) \<longrightarrow> True \<and> (True \<and> Q)"
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  2185
apply(tactic {* true1_tac @{context} 1 *})
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2186
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
  2187
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  2188
  \begin{minipage}{\textwidth}
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  2189
  @{subgoals [display]}
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  2190
  \end{minipage}*}
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2191
(*<*)oops(*>*)
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2192
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2193
text {*
148
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  2194
  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
  2195
  the conclusion according to @{ML all_true1_conv}.
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  2196
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2197
  Conversions can also be helpful for constructing meta-equality theorems.
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2198
  Such theorems are often definitions. As an example consider the following
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2199
  two ways of defining the identity function in Isabelle. 
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2200
*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2201
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2202
definition id1::"'a \<Rightarrow> 'a" 
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2203
where "id1 x \<equiv> x"
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2204
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2205
definition id2::"'a \<Rightarrow> 'a"
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2206
where "id2 \<equiv> \<lambda>x. x"
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2207
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2208
text {*
335
163ac0662211 reorganised the certified terms section; tuned
Christian Urban <urbanc@in.tum.de>
parents: 334
diff changeset
  2209
  Although both definitions define the same function, the theorems @{thm
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2210
  [source] id1_def} and @{thm [source] id2_def} are different. However it is
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2211
  easy to transform one theorem into the other. The function @{ML_ind abs_def
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2212
  in Drule} can automatically transform theorem @{thm [source] id1_def}
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2213
  into @{thm [source] id2_def} by abstracting all variables on the 
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2214
  left-hand side in the right-hand side.
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2215
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2216
  @{ML_response_fake [display,gray]
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2217
  "Drule.abs_def @{thm id1_def}"
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2218
  "id1 \<equiv> \<lambda>x. x"}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2219
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2220
  Unfortunately, Isabelle has no build-in function that transforms the
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2221
  theorems in the other direction. We can conveniently implement one using
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2222
  conversions. The feature of conversions we are using is that if we apply a
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2223
  @{ML_type cterm} to a conversion we obtain a meta-equality theorem (recall
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2224
  that the type of conversions is an abbreviation for 
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2225
  @{ML_type "cterm -> thm"}). The code of the transformation is below.
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2226
*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2227
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2228
ML %linenosgray{*fun unabs_def ctxt def =
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2229
let
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2230
  val (lhs, rhs) = Thm.dest_equals (cprop_of def)
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2231
  val xs = strip_abs_vars (term_of rhs)
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2232
  val (_, ctxt') = Variable.add_fixes (map fst xs) ctxt
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2233
  
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2234
  val thy = ProofContext.theory_of ctxt'
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2235
  val cxs = map (cterm_of thy o Free) xs
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2236
  val new_lhs = Drule.list_comb (lhs, cxs)
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2237
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2238
  fun get_conv [] = Conv.rewr_conv def
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2239
    | get_conv (_::xs) = Conv.fun_conv (get_conv xs)
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2240
in
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2241
  get_conv xs new_lhs |>  
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2242
  singleton (ProofContext.export ctxt' ctxt)
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2243
end*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2244
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2245
text {*
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2246
  In Line 3 we destruct the meta-equality into the @{ML_type cterm}s
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2247
  corresponding to the left-hand and right-hand side of the meta-equality. The
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2248
  assumption in @{ML unabs_def} is that the right-hand side is an
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2249
  abstraction. We compute the possibly empty list of abstracted variables in
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2250
  Line 4 using the function @{ML_ind strip_abs_vars}. For this we have to
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2251
  transform the @{ML_type cterm} into a @{ML_type term}. In Line 5 we
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2252
  importing these variables into the context @{ML_text ctxt'}, in order to
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2253
  export them again in Line 15.  In Line 8 we certify the list of variables,
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2254
  which in turn we apply to the @{ML_text lhs} of the definition using the
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2255
  function @{ML_ind list_comb in Drule}. In Line 11 and 12 we generate the
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2256
  conversion according to the length of the list of (abstracted) variables. If
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2257
  there are none, then we just return the conversion corresponding to the
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2258
  original definition. If there are variables, then we have to prefix this
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2259
  conversion with @{ML_ind fun_conv in Conv}. Now in Line 14 we only have to
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2260
  apply the new left-hand side to the generated conversion and obtain the the
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2261
  theorem we want to construct. As mentioned above, in Line 15 we only have to
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2262
  export the context @{ML_text ctxt'} in order to produce meta-variables for
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2263
  the theorem.  An example is as follows.
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2264
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2265
  @{ML_response_fake [display, gray]
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2266
  "unabs_def @{context} @{thm id2_def}"
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2267
  "id2 ?x1 \<equiv> ?x1"}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2268
*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2269
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2270
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2271
  To sum up this section, conversions are more general than the simplifier
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2272
  or simprocs, but you have to do more work yourself. Also conversions are
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2273
  often much less efficient than the simplifier. The advantage of conversions, 
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2274
  however, that they provide much less room for non-termination.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2275
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
  2276
  \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
  2277
  Write a tactic that does the same as the simproc in exercise
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2278
  \ref{ex:addsimproc}, but is based on conversions. You can make
166
00d153e32a53 improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
  2279
  the same assumptions as in \ref{ex:addsimproc}. 
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2280
  \end{exercise}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2281
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 170
diff changeset
  2282
  \begin{exercise}\label{ex:compare}
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2283
  Compare your solutions of Exercises~\ref{ex:addsimproc} and \ref{ex:addconversion},
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 170
diff changeset
  2284
  and try to determine which way of rewriting such terms is faster. For this you might 
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 170
diff changeset
  2285
  have to construct quite large terms. Also see Recipe \ref{rec:timing} for information 
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 170
diff changeset
  2286
  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
  2287
  \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
  2288
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2289
  \begin{readmore}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2290
  See @{ML_file "Pure/conv.ML"} for more information about conversion combinators. 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2291
  Some basic conversions are defined in  @{ML_file "Pure/thm.ML"},
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2292
  @{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
  2293
  \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
  2294
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2295
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2296
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
  2297
text {*
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
  2298
  (FIXME: check whether @{ML Pattern.match_rew} and @{ML Pattern.rewrite_term}
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
  2299
  are of any use/efficient)
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
  2300
*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2301
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
  2302
216
fcedd5bd6a35 added a declaration section (for Amine)
Christian Urban <urbanc@in.tum.de>
parents: 214
diff changeset
  2303
section {* Declarations (TBD) *}
fcedd5bd6a35 added a declaration section (for Amine)
Christian Urban <urbanc@in.tum.de>
parents: 214
diff changeset
  2304
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2305
section {* Structured Proofs (TBD) *}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2306
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2307
text {* TBD *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2308
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2309
lemma True
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2310
proof
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2311
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2312
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2313
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2314
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2315
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2316
    then have "A & B" ..
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2317
    then have C by (rule r)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2318
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2319
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2320
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2321
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2322
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2323
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2324
    note conjI [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2325
    note r [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2326
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2327
oops
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2328
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2329
ML {* 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2330
  val ctxt0 = @{context};
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2331
  val ctxt = ctxt0;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2332
  val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt;
217
75154f4d4e2f used antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
  2333
  val ([r], ctxt) = Assumption.add_assumes [@{cprop "A & B \<Longrightarrow> C"}] ctxt
75154f4d4e2f used antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
  2334
  val (this, ctxt) = Assumption.add_assumes [@{cprop "A"}, @{cprop "B"}] ctxt;
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2335
  val this = [@{thm conjI} OF this]; 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2336
  val this = r OF this;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2337
  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
  2338
  val this = Variable.export ctxt ctxt0 [this] 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2339
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2340
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2341
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
  2342
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2343
end