ProgTutorial/Tactical.thy
author Norbert Schirmer <norbert.schirmer@web.de>
Tue, 14 May 2019 11:10:53 +0200
changeset 562 daf404920ab9
parent 559 ffa5c4ec9611
child 565 cecd7a941885
permissions -rw-r--r--
Accomodate to Isabelle 2018
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Tactical
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
     2
imports Base First_Steps
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 {*
506
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
     8
   \begin{flushright}
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
     9
  {\em
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    10
  ``The first thing I would say is that when you write a program, think of\\ 
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    11
   it primarily as a work of literature. You're trying to write something\\ 
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    12
   that human beings are going to read. Don't think of it primarily as\\  
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    13
   something a computer is going to follow. The more effective you are\\ 
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    14
   at making your program readable, the more effective it's going to\\  
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    15
   be: You'll understand it today, you'll understand it next week, and\\ 
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    16
   your successors who are going to maintain and modify it will\\ 
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    17
   understand it.''}\\[1ex]
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    18
  Donald E.~Knuth, from an interview in Dr.~Dobb's Journal, 1996.
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    19
  \end{flushright}
Christian Urban <urbanc@in.tum.de>
parents: 504
diff changeset
    20
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
    21
  One of the main reason for descending to the ML-level of Isabelle is to
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
    22
  implement automatic proof procedures. Such proof procedures can
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    23
  considerably lessen the burden of manual reasoning. They are centred around
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    24
  the idea of refining a goal state using tactics. This is similar to the
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    25
  \isacommand{apply}-style reasoning at the user-level, where goals are
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    26
  modified in a sequence of proof steps until all of them are discharged.
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
    27
  In this chapter we explain how to implement simple tactics and how to combine them using
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
    28
  tactic combinators. We also describe the simplifier, simprocs and conversions.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 489
diff changeset
    31
section {* Basics of Reasoning with Tactics\label{sec:basictactics}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
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
    34
  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
    35
  into ML. Suppose the following proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    38
lemma disj_swap: 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    39
  shows "P \<or> Q \<Longrightarrow> Q \<or> P"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
apply(erule disjE)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
apply(rule disjI2)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
apply(rule disjI1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  This proof translates to the following ML-code.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
@{ML_response_fake [display,gray]
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
"let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  Goal.prove ctxt [\"P\", \"Q\"] [] goal 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    56
   (fn _ =>  eresolve_tac @{context} [@{thm disjE}] 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    57
             THEN resolve_tac @{context} [@{thm disjI2}] 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    58
             THEN assume_tac @{context} 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    59
             THEN resolve_tac @{context} [@{thm disjI1}] 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    60
             THEN assume_tac @{context} 1)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    63
  To start the proof, the function @{ML_ind prove in Goal} sets up a goal
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    64
  state for proving the goal @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. We can give this
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    65
  function some assumptions in the third argument (there are no assumption in
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    66
  the proof at hand). The second argument stands for a list of variables
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    67
  (given as strings). This list contains the variables that will be turned
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    68
  into schematic variables once the goal is proved (in our case @{text P} and
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    69
  @{text Q}). The last argument is the tactic that proves the goal. This
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
    70
  tactic can make use of the local assumptions (there are none in this
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    71
  example). The tactics @{ML_ind eresolve_tac in Tactic}, @{ML_ind resolve_tac in Tactic} and
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    72
  @{ML_ind assume_tac in Tactic} in the code above correspond roughly to @{text
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
    73
  erule}, @{text rule} and @{text assumption}, respectively. The combinator
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
    74
  @{ML THEN} strings the tactics together.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    76
  TBD: Write something about @{ML_ind prove_common in Goal}.
437
Christian Urban <urbanc@in.tum.de>
parents: 436
diff changeset
    77
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  \begin{readmore}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    79
  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
    80
  \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
    81
  "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
    82
  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
    83
  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
    84
  Manual.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
436
373f99b1221a added something about raw_tactic
Christian Urban <urbanc@in.tum.de>
parents: 424
diff changeset
    88
  \index{tactic@ {\tt\slshape{}tactic}}
373f99b1221a added something about raw_tactic
Christian Urban <urbanc@in.tum.de>
parents: 424
diff changeset
    89
  \index{raw\_tactic@ {\tt\slshape{}raw\_tactic}}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
    90
  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
    91
  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
    92
  the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  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
    94
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    97
ML %grayML{*fun foo_tac ctxt = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    98
      (eresolve_tac ctxt [@{thm disjE}] 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    99
       THEN resolve_tac  ctxt [@{thm disjI2}] 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   100
       THEN assume_tac  ctxt 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   101
       THEN resolve_tac ctxt [@{thm disjI1}] 1
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   102
       THEN assume_tac  ctxt 1)*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
text {* and the Isabelle proof: *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   106
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   107
  shows "P \<or> Q \<Longrightarrow> Q \<or> P"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   108
apply(tactic {* foo_tac @{context} *})
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   112
  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
   113
  user-level of Isabelle the tactic @{ML foo_tac} or 
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
   114
  any other function that returns a tactic. There are some goal transformations
436
373f99b1221a added something about raw_tactic
Christian Urban <urbanc@in.tum.de>
parents: 424
diff changeset
   115
  that are performed by @{text "tactic"}. They can be avoided by using 
373f99b1221a added something about raw_tactic
Christian Urban <urbanc@in.tum.de>
parents: 424
diff changeset
   116
  @{text "raw_tactic"} instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
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
   118
  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
   119
  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
   120
  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
   121
  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
   122
  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
   123
  you can write
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   126
ML %grayML{*fun foo_tac' ctxt = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   127
      (eresolve_tac ctxt [@{thm disjE}] 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   128
       THEN' resolve_tac ctxt [@{thm disjI2}] 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   129
       THEN' assume_tac ctxt 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   130
       THEN' resolve_tac ctxt [@{thm disjI1}] 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   131
       THEN' assume_tac ctxt)*}text_raw{*\label{tac:footacprime}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
text {* 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   134
  where @{ML_ind THEN' in Tactical} is used instead of @{ML THEN in
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   135
  Tactical}. (For most combinators that combine tactics---@{ML THEN} is only
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   136
  one such combinator---a ``primed'' version exists.)  With @{ML foo_tac'} you
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   137
  can give the number for the subgoal explicitly when the tactic is called. So
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   138
  in the next proof you can first discharge the second subgoal, and
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
   139
  then the first.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   142
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   143
  shows "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   144
  and   "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   145
apply(tactic {* foo_tac' @{context} 2 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   146
apply(tactic {* foo_tac' @{context} 1 *})
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
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
   150
  This kind of addressing is more difficult to achieve when the goal is 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   151
  hard-coded inside the tactic. 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   152
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
   153
  The tactics @{ML foo_tac} and @{ML foo_tac'} are specific
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
   154
  to goals of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not of
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   155
  this form, then these tactics return the error message:\footnote{To be
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   156
  precise, tactics do not produce this error message; the message originates from the
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   157
  \isacommand{apply} wrapper that calls the tactic.}
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   158
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   159
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   160
  \begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   161
  @{text "*** empty result sequence -- proof command failed"}\\
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   162
  @{text "*** At command \"apply\"."}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   163
  \end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   164
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   165
  This means they failed. The reason for this error message is that tactics
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   166
  are functions mapping a goal state to a (lazy) sequence of successor
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   167
  states. Hence the type of a tactic is:
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   168
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
   170
ML %grayML{*type tactic = thm -> thm Seq.seq*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   172
text {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   173
  By convention, if a tactic fails, then it should return the empty sequence. 
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
   174
  Therefore, your own tactics should not raise exceptions 
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 556
diff changeset
   175
  willy-nilly; only in very grave failure situations should a tactic raise the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   176
  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
   177
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   178
  The simplest tactics are @{ML_ind no_tac in Tactical} and 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   179
  @{ML_ind all_tac in Tactical}. The first returns the empty sequence and 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   180
  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
   181
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   182
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
   183
ML %grayML{*fun no_tac thm = Seq.empty*}
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
   184
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   185
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
   186
  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
   187
  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
   188
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   189
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
   190
ML %grayML{*fun all_tac thm = Seq.single 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
   191
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   192
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   193
  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
   194
  with the proof.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   196
  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
   197
  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
   198
  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
   199
  @{ML foo_tac'}: either using the first assumption or the second.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   202
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   203
  shows "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   204
apply(tactic {* foo_tac' @{context} 1 *})
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
back
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
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
   208
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   210
  By using \isacommand{back}, we construct the proof that uses the
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   211
  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
   212
  assumption is used, in more interesting cases provability might depend
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   213
  on exploring different possibilities.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   214
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  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
   217
  sequences. In day-to-day Isabelle programming, however, one rarely 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   218
  constructs sequences explicitly, but uses the predefined tactics and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   219
  tactic combinators instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   222
  It might be surprising that tactics, which transform
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   223
  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
   224
  (sequences). The surprise resolves by knowing that every 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   225
  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
   226
  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
   227
  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
   228
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   229
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
   230
ML %grayML{*fun my_print_tac ctxt thm =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   231
let
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   232
  val _ = tracing (Pretty.string_of (pretty_thm_no_vars ctxt thm))
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   233
in 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   234
  Seq.single thm
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   235
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
   236
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   237
text {*
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   238
  which prints out the given theorem (using the string-function defined in
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   239
  Section~\ref{sec:printing}) and then behaves like @{ML all_tac}. With this
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   240
  tactic we are in the position to inspect every goal state in a proof. For
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   241
  this consider the proof in Figure~\ref{fig:goalstates}: as can be seen,
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   242
  internally every goal state is an implication of the form
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   243
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   244
  @{text[display] "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C"}
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   245
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   246
  where @{term C} is the goal to be proved and the @{term "A\<^sub>i"} are
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   247
  the subgoals. So after setting up the lemma, the goal state is always of the
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   248
  form @{text "C \<Longrightarrow> #C"}; when the proof is finished we are left with @{text
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   249
  "#C"}. Since the goal @{term C} can potentially be an implication, there is a
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   250
  ``protector'' wrapped around it (the wrapper is the outermost constant
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 552
diff changeset
   251
  @{text "Const (\"Pure.prop\", bool \<Rightarrow> bool)"}; in the figure we make it visible
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   252
  as a @{text #}). This wrapper prevents that premises of @{text C} are
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   253
  misinterpreted as open subgoals. While tactics can operate on the subgoals
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   254
  (the @{text "A\<^sub>i"} above), they are expected to leave the conclusion
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   255
  @{term C} intact, with the exception of possibly instantiating schematic
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   256
  variables. This instantiations of schematic variables can be observed 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   257
  on the user level. Have a look at the following definition and proof.
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   258
*}
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   259
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   260
text_raw {*
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   261
  \begin{figure}[p]
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   262
  \begin{boxedminipage}{\textwidth}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   263
  \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
   264
*}
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 552
diff changeset
   265
notation (output) "Pure.prop"  ("#_" [1000] 1000)
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
   266
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   267
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   268
  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
   269
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
   270
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   271
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
   272
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   273
      \end{minipage}\medskip      
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   274
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   275
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   276
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   277
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   278
      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
   279
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   280
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   281
      \end{minipage}\medskip
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
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
   284
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
   285
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
   286
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   287
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
   288
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   289
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   290
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   291
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   292
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   293
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   294
      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
   295
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   296
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   297
      \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
   298
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   299
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   300
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
   301
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
   302
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   303
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
   304
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   305
      \end{minipage}\medskip
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   306
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   307
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   308
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   309
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   310
      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
   311
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   312
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   313
      \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
   314
*}
5e309df58557 general 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
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   316
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
   317
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
   318
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   319
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
   320
      @{subgoals [display]} 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   321
      \end{minipage}\medskip 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   322
  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   323
      \begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   324
      \small\colorbox{gray!20}{
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   325
      \begin{tabular}{@ {}l@ {}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   326
      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
   327
      @{raw_goal_state}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   328
      \end{tabular}}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   329
      \end{minipage}\medskip   
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   330
   *}
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
   331
(*<*)oops(*>*)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   332
text_raw {*  
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   333
  \end{isabelle}
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   334
  \end{boxedminipage}
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   335
  \caption{The figure shows an Isabelle snippet of a proof where each
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   336
  intermediate goal state is printed by the Isabelle system and by @{ML
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   337
  my_print_tac}. The latter shows the goal state as represented internally
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   338
  (highlighted boxes). This tactic shows that every goal state in Isabelle is
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   339
  represented by a theorem: when you start the proof of \mbox{@{text "\<lbrakk>A; B\<rbrakk> \<Longrightarrow>
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   340
  A \<and> B"}} the theorem is @{text "(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B) \<Longrightarrow> #(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B)"}; when
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   341
  you finish the proof the theorem is @{text "#(\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and>
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   342
  B)"}.\label{fig:goalstates}}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   343
  \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
   344
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   345
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   346
definition 
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   347
  EQ_TRUE 
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   348
where
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   349
  "EQ_TRUE X \<equiv> (X = True)"
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   350
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   351
schematic_goal test: 
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   352
  shows "EQ_TRUE ?X"
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   353
unfolding EQ_TRUE_def
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   354
by (rule refl)
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   355
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   356
text {*
422
e79563b14b0f updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   357
  By using \isacommand{schematic\_lemma} it is possible to prove lemmas including
e79563b14b0f updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   358
  meta-variables on the user level. However, the proved theorem is not @{text "EQ_TRUE ?X"}, 
e79563b14b0f updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   359
  as one might expect, but @{thm test}. We can test this with:
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   360
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   361
  \begin{isabelle}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   362
  \isacommand{thm}~@{thm [source] test}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   363
  @{text ">"}~@{thm test}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 351
diff changeset
   364
  \end{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
   365
 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   366
  The reason for this result is that the schematic variable @{text "?X"} 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   367
  is instantiated inside the proof to be @{term True} and then the 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   368
  instantiation propagated to the ``outside''.
359
be6538c7b41d polished
Christian Urban <urbanc@in.tum.de>
parents: 358
diff changeset
   369
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
   370
  \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
   371
  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
   372
  \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
   373
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
   374
*}
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 99
diff changeset
   375
422
e79563b14b0f updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 418
diff changeset
   376
194
8cd51a25a7ca some polishing
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   377
section {* Simple Tactics\label{sec:simpletacs} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   379
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   380
  In this section we will introduce more of the simple tactics in Isabelle. The 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   381
  first one is @{ML_ind print_tac in Tactical}, which is quite useful 
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   382
  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
   383
  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
   384
  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
   385
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   386
 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   387
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   388
  shows "False \<Longrightarrow> True"
556
3c214b215f7e some small updates for Isabelle and corrections in the Parsing chapter
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
   389
apply(tactic {* print_tac @{context} "foo message" *})
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   390
txt{*gives:
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   391
     \begin{isabelle}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   392
     @{text "foo message"}\\[3mm] 
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   393
     @{prop "False \<Longrightarrow> True"}\\
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   394
     @{text " 1. False \<Longrightarrow> True"}\\
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   395
     \end{isabelle}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   396
   *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   397
(*<*)oops(*>*)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   398
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   399
text {*
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   400
  A simple tactic for easy discharge of any proof obligations, even difficult ones, is 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   401
  @{ML_ind cheat_tac in Skip_Proof} in the structure @{ML_struct Skip_Proof}. 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   402
  This tactic corresponds to the Isabelle command \isacommand{sorry} and is 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   403
  sometimes useful during the development of tactics.
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   404
*}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   405
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   406
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   407
  shows "False" and "Goldbach_conjecture"  
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   408
apply(tactic {* Skip_Proof.cheat_tac @{context} 1 *})
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   409
txt{*\begin{minipage}{\textwidth}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   410
     @{subgoals [display]}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   411
     \end{minipage}*}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   412
(*<*)oops(*>*)
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   413
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   414
text {*
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   415
  Another simple tactic is the function @{ML_ind assume_tac in Tactic}, which, as shown 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   416
  earlier, corresponds to the assumption method.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   417
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   418
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   419
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   420
  shows "P \<Longrightarrow> P"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   421
apply(tactic {* assume_tac @{context} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   422
txt{*\begin{minipage}{\textwidth}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   423
     @{subgoals [display]}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   424
     \end{minipage}*}
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   425
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   427
text {*
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   428
  Similarly, @{ML_ind resolve_tac in Tactic}, @{ML_ind dresolve_tac in Tactic}, @{ML_ind eresolve_tac
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   429
  in Tactic} and @{ML_ind forward_tac in Tactic} correspond (roughly) to @{text
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   430
  rule}, @{text drule}, @{text erule} and @{text frule}, respectively. Each of
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   431
  them takes a theorem as argument and attempts to apply it to a goal. Below
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   432
  are three self-explanatory examples.
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   433
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   434
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   435
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   436
  shows "P \<and> Q"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   437
apply(tactic {* resolve_tac @{context} [@{thm conjI}] 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   438
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   439
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   440
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   443
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   444
  shows "P \<and> Q \<Longrightarrow> False"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   445
apply(tactic {* eresolve_tac @{context} [@{thm conjE}] 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   446
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   447
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   448
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   451
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   452
  shows "False \<and> True \<Longrightarrow> False"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   453
apply(tactic {* dresolve_tac @{context} [@{thm conjunct2}] 1 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   454
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   455
     @{subgoals [display]}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   456
     \end{minipage}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
text {*
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   460
  The function @{ML_ind resolve_tac in Tactic} 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   461
  expects a list of theorems as argument. From this list it will apply the
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   462
  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
   463
  explored via the lazy sequences mechanism). Given the code
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   466
ML %grayML{*fun resolve_xmp_tac ctxt = resolve_tac ctxt [@{thm impI}, @{thm conjI}]*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   467
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   468
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   469
  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
   470
  implication is analysed and then an outermost conjunction.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   471
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   472
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   473
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   474
  shows "C \<longrightarrow> (A \<and> B)" 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   475
  and   "(A \<longrightarrow> B) \<and> C"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   476
apply(tactic {* resolve_xmp_tac @{context} 1 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   477
apply(tactic {* resolve_xmp_tac @{context} 2 *})
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   478
txt{*\begin{minipage}{\textwidth}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   479
     @{subgoals [display]} 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   480
     \end{minipage}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   481
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   482
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   483
text {* 
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
   484
  Similar versions taking a list of theorems exist for the tactics 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   485
   @{ML_ind dresolve_tac in Tactic},  
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   486
  @{ML_ind eresolve_tac in Tactic} and so on.
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   487
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   488
  Another simple tactic is @{ML_ind cut_facts_tac in Tactic}. It inserts a
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   489
  list of theorems into the assumptions of the current goal state. Below we
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   490
  will insert the definitions for the constants @{term True} and @{term
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   491
  False}. So
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   492
*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   493
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   494
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   495
  shows "True \<noteq> False"
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   496
apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   497
txt{*produces the goal state
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   498
     \begin{isabelle}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   499
     @{subgoals [display]} 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   500
     \end{isabelle}*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   501
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   502
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   503
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   504
  Often proofs on the ML-level involve elaborate operations on assumptions and 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   505
  @{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
   506
  shown so far is very unwieldy and brittle. Some convenience and
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   507
  safety is provided by the functions @{ML_ind FOCUS in Subgoal} and 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   508
  @{ML_ind SUBPROOF in Subgoal}. These tactics fix the parameters 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   509
  and bind the various components of a goal state to a record. 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   510
  To see what happens, suppose the function defined in Figure~\ref{fig:sptac}, which
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   511
  takes a record and just prints out the contents of this record. Then consider
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   512
  the proof:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   513
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   514
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   515
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   516
text_raw{*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
   517
\begin{figure}[t]
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
   518
\begin{minipage}{\textwidth}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   519
\begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   520
*}
294
ee9d53fbb56b made changes for SUBPROOF and sat_tac
Christian Urban <urbanc@in.tum.de>
parents: 292
diff changeset
   521
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   522
ML %grayML{*fun foc_tac {context, params, prems, asms, concl, schematics} = 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   523
let 
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   524
  fun assgn1 f1 f2 xs =
541
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   525
    let
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   526
      val out = map (fn (x, y) => Pretty.enum ":=" "" "" [f1 x, f2 y]) xs
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   527
    in
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   528
      Pretty.list "" "" out
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   529
    end 
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   530
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   531
  fun assgn2 f xs = assgn1 (fn (n,T) => pretty_term context (Var (n,T))) f  xs 
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   532
 
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   533
  val pps = map (fn (s, pp) => Pretty.block [Pretty.str s, pp])
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   534
    [("params: ", assgn1 Pretty.str (pretty_cterm context) params),
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   535
     ("assumptions: ", pretty_cterms context asms),
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   536
     ("conclusion: ", pretty_cterm context concl),
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   537
     ("premises: ", pretty_thms_no_vars context prems),   
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   538
     ("schematics: ", assgn2 (pretty_cterm context) (snd schematics))]
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   539
in
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   540
  tracing (Pretty.string_of (Pretty.chunks pps)); all_tac 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   541
end*}
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 298
diff changeset
   542
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   543
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   544
\end{isabelle}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
   545
\end{minipage}
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   546
\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
   547
 @{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
   548
  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
   549
  and @{ML_type thm}s.\label{fig:sptac}}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   550
\end{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   551
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   552
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   553
schematic_goal
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   554
  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
   555
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
   556
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   557
txt {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   558
  The tactic produces the following printout:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   559
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   560
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   561
  \begin{tabular}{ll}
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   562
  params:      & @{text "x:= x"}, @{text "y:= y"}\\
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
   563
  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
   564
  conclusion:  & @{term "B y x \<longrightarrow> C (z y) x"}\\
541
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   565
  premises:    & @{term "A x y"}\\
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   566
  schematics:  & @{text "?z:=z"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   567
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   568
  \end{quote}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   569
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   570
  The @{text params} and @{text schematics} stand for list of pairs where the
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   571
  left-hand side of @{text ":="} is replaced by the right-hand side inside the
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   572
  tactic.  Notice that in the actual output the variables @{term x} and @{term
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   573
  y} have a brown colour. Although they are parameters in the original goal,
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   574
  they are fixed inside the tactic. By convention these fixed variables are
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   575
  printed in brown colour.  Similarly the schematic variable @{text ?z}. The
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   576
  assumption, or premise, @{prop "A x y"} is bound as @{ML_type cterm} to the
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   577
  record-variable @{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
   578
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   579
  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
   580
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   581
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   582
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
   583
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
   584
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   585
txt {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   586
  then the tactic prints out: 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   587
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   588
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   589
  \begin{tabular}{ll}
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
   590
  params:      & @{text "x:= x"}, @{text "y:= y"}\\
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
   591
  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
   592
  conclusion:  & @{term "C (z y) x"}\\
541
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   593
  premises:    & @{term "A x y"}, @{term "B y x"}\\
96d10631eec2 added readme and fixed output in Subgoal.FOCUS section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   594
  schematics:  & @{text "?z:=z"}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   595
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   596
  \end{quote}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   597
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   598
(*<*)oops(*>*)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   599
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   600
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   601
  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
   602
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   603
  One difference between the tactics @{ML SUBPROOF} and @{ML FOCUS in Subgoal}
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   604
  is that the former expects that the goal is solved completely, which the
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   605
  latter does not. Another is that @{ML SUBPROOF} cannot instantiate any schematic
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   606
  variables.
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   607
411
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   608
  Observe that inside @{ML FOCUS in Subgoal} and @{ML SUBPROOF}, we are in a goal
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   609
  state where there is only a conclusion. This means the tactics @{ML dresolve_tac} and 
411
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   610
  the like are of no use for manipulating the goal state. The assumptions inside 
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   611
  @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are given as cterms and theorems in 
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   612
  the arguments @{text "asms"} and @{text "prems"}, respectively. This 
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   613
  means we can apply them using the usual assumption tactics. With this you can 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   614
  for example easily implement a tactic that behaves almost like @{ML assume_tac}:
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   615
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   616
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   617
ML %grayML{*val atac' = Subgoal.FOCUS (fn {prems, context, ...} => resolve_tac context prems 1)*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   618
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   619
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   620
  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
   621
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   622
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   623
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   624
  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
   625
apply(tactic {* atac' @{context} 1 *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   626
txt{* it will produce
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   627
      @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   628
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   629
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   630
text {*
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   631
  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
   632
  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
   633
  @{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
   634
  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
   635
  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
   636
  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
   637
  goal by just writing:
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   638
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   639
*}
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   640
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   641
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   642
  shows "True" 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   643
  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
   644
apply(tactic {* atac' @{context} 2 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   645
apply(rule TrueI)
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   646
done
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   647
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
text {*
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 449
diff changeset
   649
  To sum up, both @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are rather
411
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   650
  convenient, but can impose a considerable run-time penalty in automatic
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   651
  proofs. If speed is of the essence, then maybe the two lower level combinators 
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   652
  described next are more appropriate.
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   653
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   654
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   655
  \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
   656
  The functions @{ML FOCUS in Subgoal} and @{ML SUBPROOF} are defined in 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   657
  @{ML_file "Pure/Isar/subgoal.ML"} and also described in 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   658
  \isccite{sec:results}. 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   659
  \end{readmore}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   660
411
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   661
  Similar but less powerful functions than @{ML FOCUS in Subgoal},
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   662
  respectively @{ML SUBPROOF}, are @{ML_ind SUBGOAL in Tactical} and @{ML_ind
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   663
  CSUBGOAL in Tactical}. They allow you to 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
   664
  presents the subgoal as a @{ML_type term}, while the latter as a @{ML_type
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   665
  cterm}). With them you can implement a tactic that applies a rule according
151
7e0bf13bf743 added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 150
diff changeset
   666
  to the topmost logic connective in the subgoal (to illustrate this we only
411
24fc00319c4a polised the section about Subgoal.FOCUS
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   667
  analyse a few connectives). The code of the tactic is as follows.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   668
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   670
ML %linenosgray{*fun select_tac ctxt (t, i) =
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   671
  case t of
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   672
     @{term "Trueprop"} $ t' => select_tac ctxt (t', i)
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   673
   | @{term "(\<Longrightarrow>)"} $ _ $ t' => select_tac ctxt (t', i)
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   674
   | @{term "(\<and>)"} $ _ $ _ => resolve_tac ctxt [@{thm conjI}] i
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   675
   | @{term "(\<longrightarrow>)"} $ _ $ _ => resolve_tac ctxt [@{thm impI}] i
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   676
   | @{term "Not"} $ _ => resolve_tac ctxt [@{thm notI}] i
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   677
   | Const (@{const_name "All"}, _) $ _ => resolve_tac ctxt [@{thm allI}] i
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
   678
   | _ => all_tac*}text_raw{*\label{tac:selecttac}*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   679
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   680
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   681
  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
   682
  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
   683
  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
   684
  analyse. Otherwise goals like @{prop "A \<and> B"} are not properly
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   685
  analysed. Similarly with meta-implications in the next line.  While for the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   686
  first five patterns we can use the @{text "@term"}-antiquotation to
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   687
  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
   688
  way. The reason is that an antiquotation would fix the type of the
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   689
  quantified variable. So you really have to construct this pattern using the
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   690
  basic term-constructors. This is not necessary in the other cases, because their
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   691
  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
   692
  bool}. (See Section \ref{sec:terms_types_manually} about constructing terms
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   693
  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
   694
  Consequently, @{ML select_tac} never fails.
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   695
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   696
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   697
  Let us now see how to apply this tactic. Consider the four goals:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   698
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   699
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   700
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   701
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   702
  shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   703
apply(tactic {* SUBGOAL (select_tac @{context}) 4 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   704
apply(tactic {* SUBGOAL (select_tac @{context}) 3 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   705
apply(tactic {* SUBGOAL (select_tac @{context}) 2 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   706
apply(tactic {* SUBGOAL (select_tac @{context}) 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   707
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   708
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   709
      \end{minipage} *}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   710
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   711
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   712
text {*
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   713
  where in all but the last the tactic applies an introduction rule. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   714
  Note that we applied the tactic to the goals in ``reverse'' order. 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   715
  This is a trick in order to be independent from the subgoals that are 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   716
  produced by the rule. If we had applied it in the other order 
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   717
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   718
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   719
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   720
  shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   721
apply(tactic {* SUBGOAL (select_tac @{context}) 1 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   722
apply(tactic {* SUBGOAL (select_tac @{context}) 3 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   723
apply(tactic {* SUBGOAL (select_tac @{context}) 4 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   724
apply(tactic {* SUBGOAL (select_tac @{context}) 5 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   725
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   726
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   727
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   728
  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
   729
  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
   730
  the ``reverse application'' is easy to implement.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   731
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
   732
  Of course, this example is
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   733
  contrived: there are much simpler methods available in Isabelle for
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
   734
  implementing a tactic analysing a goal according to its topmost
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
   735
  connective. These simpler methods use tactic combinators, which we will
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   736
  explain in the next section. But before that we will show how
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   737
  tactic application can be constrained.
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   738
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   739
  \begin{readmore}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   740
  The functions @{ML SUBGOAL} and @{ML CSUBGOAL} are defined in @{ML_file "Pure/tactical.ML"}.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   741
  \end{readmore}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   742
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   743
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   744
  Since @{ML_ind resolve_tac in Tactic} and the like use higher-order unification, an
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   745
  automatic proof procedure based on them might become too fragile, if it just
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   746
  applies theorems as shown above. The reason is that a number of theorems
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   747
  introduce schematic variables into the goal state. Consider for example the
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   748
  proof
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   749
*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   750
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   751
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   752
  shows "\<forall>x \<in> A. P x \<Longrightarrow> Q x"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   753
apply(tactic {* dresolve_tac @{context} [@{thm bspec}] 1 *})
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   754
txt{*\begin{minipage}{\textwidth}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   755
     @{subgoals [display]} 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   756
     \end{minipage}*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   757
(*<*)oops(*>*)
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   758
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   759
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   760
  where the application of theorem @{text bspec} generates two subgoals involving the
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   761
  new schematic variable @{text "?x"}. Now, if you are not careful, tactics 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   762
  applied to the first subgoal might instantiate this schematic variable in such a 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   763
  way that the second subgoal becomes unprovable. If it is clear what the @{text "?x"}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   764
  should be, then this situation can be avoided by introducing a more
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   765
  constrained version of the @{text bspec}-theorem. One way to give such 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   766
  constraints is by pre-instantiating theorems with other theorems. 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   767
  The function @{ML_ind RS in Drule}, for example, does this.
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   768
  
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   769
  @{ML_response_fake [display,gray]
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   770
  "@{thm disjI1} RS @{thm conjI}" "\<lbrakk>?P1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q"}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   771
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   772
  In this example it instantiates the first premise of the @{text conjI}-theorem 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   773
  with the theorem @{text disjI1}. If the instantiation is impossible, as in the 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   774
  case of
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   775
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   776
  @{ML_response_fake_both [display,gray]
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   777
  "@{thm conjI} RS @{thm mp}" 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   778
"*** Exception- THM (\"RSN: no unifiers\", 1, 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   779
[\"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\", \"\<lbrakk>?P \<longrightarrow> ?Q; ?P\<rbrakk> \<Longrightarrow> ?Q\"]) raised"}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   780
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   781
  then the function raises an exception. The function @{ML_ind  RSN in Drule} 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   782
  is similar to @{ML RS}, but takes an additional number as argument. This 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   783
  number makes explicit which premise should be instantiated. 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   784
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   785
  If you want to instantiate more than one premise of a theorem, you can use 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   786
  the function @{ML_ind MRS in Drule}:
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   787
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   788
  @{ML_response_fake [display,gray]
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 465
diff changeset
   789
  "[@{thm disjI1}, @{thm disjI2}] MRS @{thm conjI}" 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 465
diff changeset
   790
  "\<lbrakk>?P2; ?Q1\<rbrakk> \<Longrightarrow> (?P2 \<or> ?Q2) \<and> (?P1 \<or> ?Q1)"}
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   791
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   792
  If you need to instantiate lists of theorems, you can use the
510
500d1abbc98c updated to Isabelle 15 Feb
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
   793
  functions @{ML_ind RL in Drule} and @{ML_ind OF in Drule}. For 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   794
  example in the code below, every theorem in the second list is 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   795
  instantiated with every theorem in the first.
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   796
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   797
  @{ML_response_fake [display,gray]
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   798
"let
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   799
  val list1 = [@{thm impI}, @{thm disjI2}]
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   800
  val list2 = [@{thm conjI}, @{thm disjI1}]
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   801
in
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 465
diff changeset
   802
  list1 RL list2
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   803
end" 
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 465
diff changeset
   804
"[\<lbrakk>?P1 \<Longrightarrow> ?Q1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<longrightarrow> ?Q1) \<and> ?Q, 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 465
diff changeset
   805
 \<lbrakk>?Q1; ?Q\<rbrakk> \<Longrightarrow> (?P1 \<or> ?Q1) \<and> ?Q,
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 465
diff changeset
   806
 (?P1 \<Longrightarrow> ?Q1) \<Longrightarrow> (?P1 \<longrightarrow> ?Q1) \<or> ?Q, 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 465
diff changeset
   807
 ?Q1 \<Longrightarrow> (?P1 \<or> ?Q1) \<or> ?Q]"}
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   808
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   809
  \begin{readmore}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   810
  The combinators for instantiating theorems with other theorems are 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   811
  defined in @{ML_file "Pure/drule.ML"}.
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   812
  \end{readmore}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   813
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   814
  Higher-order unification is also often in the way when applying certain 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   815
  congruence theorems. For example we would expect that the theorem 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   816
  @{thm [source] cong}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   817
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   818
  \begin{isabelle}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   819
  \isacommand{thm}~@{thm [source] cong}\\
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   820
  @{text "> "}~@{thm cong}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   821
  \end{isabelle}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   822
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   823
  is applicable in the following proof producing the subgoals
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   824
  @{term "t x = s u"} and @{term "y = w"}. 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   825
*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   826
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   827
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   828
  fixes x y u w::"'a"
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   829
  shows "t x y = s u w"
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   830
apply(rule cong)
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   831
txt{*\begin{minipage}{\textwidth}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   832
     @{subgoals [display]}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   833
     \end{minipage}*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   834
(*<*)oops(*>*)
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   835
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   836
text {*
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   837
  As you can see this is unfortunately \emph{not} the case if we apply @{thm [source] 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   838
  cong} with the method @{text rule}. The problem is
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   839
  that higher-order unification produces an instantiation that is not the
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   840
  intended one. While we can use \isacommand{back} to interactively find the
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   841
  intended instantiation, this is not an option for an automatic proof
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   842
  procedure. Fortunately, the tactic @{ML_ind cong_tac in Cong_Tac} helps 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   843
  with applying congruence theorems and finding the intended instantiation.
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   844
  For example
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   845
*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   846
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   847
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   848
  fixes x y u w::"'a"
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   849
  shows "t x y = s u w"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   850
apply(tactic {* Cong_Tac.cong_tac @{context} @{thm cong} 1 *})
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   851
txt{*\begin{minipage}{\textwidth}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   852
     @{subgoals [display]}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   853
     \end{minipage}*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   854
(*<*)oops(*>*)
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   855
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   856
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   857
  However, frequently it is necessary to explicitly match a theorem against a goal
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   858
  state and in doing so construct manually an appropriate instantiation. Imagine 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   859
  you have the theorem
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   860
*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   861
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   862
lemma rule:
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   863
  shows "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> R (f x) (g y)"
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   864
sorry
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   865
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   866
text {* 
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   867
  and you want to apply it to the goal @{term "(t\<^sub>1 t\<^sub>2) \<le>
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   868
  (s\<^sub>1 s\<^sub>2)"}. Since in the theorem all variables are
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   869
  schematic, we have a nasty higher-order unification problem and @{text rtac}
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   870
  will not be able to apply this rule in the way we want. For the goal at hand 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   871
  we want to use it so that @{term R} is instantiated to the constant @{text \<le>} and
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   872
  similarly ``obvious'' instantiations for the other variables.  To achieve this 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   873
  we need to match the conclusion of 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   874
  @{thm [source] rule} against the goal reading off an instantiation for
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   875
  @{thm [source] rule}. For this the function @{ML_ind first_order_match in Thm} 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   876
  matches two @{ML_type cterm}s and produces, in the successful case, a matcher 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   877
  that can be used to instantiate the theorem. The instantiation 
465
886a7c614ced updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 458
diff changeset
   878
  can be done with the function @{ML_ind instantiate_normalize in Drule}. To automate 
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   879
  this we implement the following function.
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   880
*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   881
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   882
ML %linenosgray{*fun fo_rtac thm = Subgoal.FOCUS (fn {context, concl, ...} =>
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   883
  let 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   884
    val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm)
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   885
    val insts = Thm.first_order_match (concl_pat, concl)
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   886
  in
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   887
    resolve_tac context [(Drule.instantiate_normalize insts thm)] 1
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   888
  end)*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   889
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   890
text {*
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   891
  Note that we use @{ML FOCUS in Subgoal} because it gives us directly access
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   892
  to the conclusion of the goal state, but also because this function 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   893
  takes care of correctly handling parameters that might be present
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   894
  in the goal state. In Line 3 we use the function @{ML_ind strip_imp_concl in Drule}
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   895
  for calculating the conclusion of a theorem (produced as @{ML_type cterm}).
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   896
  An example of @{ML fo_rtac} is as follows.
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   897
*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   898
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   899
lemma
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   900
  shows "\<And>t\<^sub>1 s\<^sub>1 (t\<^sub>2::'a) (s\<^sub>2::'a). (t\<^sub>1 t\<^sub>2) \<le> (s\<^sub>1 s\<^sub>2)"
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   901
apply(tactic {* fo_rtac @{thm rule} @{context} 1 *})
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   902
txt{*\begin{minipage}{\textwidth}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   903
     @{subgoals [display]}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   904
     \end{minipage}*}
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   905
(*<*)oops(*>*)
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   906
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   907
text {*
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   908
  We obtain the intended subgoals and also the parameters are correctly
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   909
  introduced in both of them. Such manual instantiations are quite frequently
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   910
  necessary in order to appropriately constrain the application of theorems. 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   911
  Otherwise one can end up with ``wild'' higher-order unification  problems 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   912
  that make automatic proofs fail.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   913
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   914
  \begin{readmore}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   915
  Functions for matching @{ML_type cterm}s are defined in @{ML_file "Pure/thm.ML"}.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   916
  Functions for instantiating schematic variables in theorems are defined
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   917
  in @{ML_file "Pure/drule.ML"}.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   918
  \end{readmore}
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
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   921
section {* Tactic Combinators *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   922
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   923
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   924
  The purpose of tactic combinators is to build compound tactics out of
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   925
  smaller tactics. In the previous section we already used @{ML_ind THEN in Tactical}, 
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   926
  which just strings together two tactics in a sequence. For example:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   927
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   928
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   929
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   930
  shows "(Foo \<and> Bar) \<and> False"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   931
apply(tactic {* resolve_tac @{context} [@{thm conjI}] 1 THEN resolve_tac @{context} [@{thm conjI}] 1 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   932
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   933
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   934
       \end{minipage} *}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   935
(*<*)oops(*>*)
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   936
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   937
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   938
  If you want to avoid the hard-coded subgoal addressing in each component, 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   939
  then, as seen earlier, you can use the ``primed'' version of @{ML THEN}. 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
   940
  For example:
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   941
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   942
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   943
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   944
  shows "(Foo \<and> Bar) \<and> False"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   945
apply(tactic {* (resolve_tac @{context} [@{thm conjI}] THEN' resolve_tac @{context} [@{thm conjI}]) 1 *})
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   946
txt {* \begin{minipage}{\textwidth}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   947
       @{subgoals [display]} 
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   948
       \end{minipage} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   949
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   950
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   951
text {* 
404
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   952
  Here you have to specify the subgoal of interest only once and it is
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   953
  consistently applied to the component.  For most tactic combinators such a
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   954
  ``primed'' version exists and in what follows we will usually prefer it over
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   955
  the ``unprimed'' one.
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   956
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   957
  The tactic combinator @{ML_ind RANGE in Tactical} takes a list of @{text n}
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   958
  tactics, say, and applies them to each of the first @{text n} subgoals. 
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   959
  For example below we first apply the introduction rule for conjunctions
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   960
  and then apply a tactic to each of the two subgoals. 
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   961
*}
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   962
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   963
lemma 
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   964
  shows "A \<Longrightarrow> True \<and> A"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   965
apply(tactic {* (resolve_tac @{context} [@{thm conjI}] 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   966
                 THEN' RANGE [resolve_tac @{context} [@{thm TrueI}], assume_tac @{context}]) 1 *})
404
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   967
txt {* \begin{minipage}{\textwidth}
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   968
       @{subgoals [display]} 
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   969
       \end{minipage} *}
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   970
(*<*)oops(*>*)
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   971
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   972
text {*
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   973
  If there is a list of tactics that should all be tried out in sequence on
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   974
  one specified subgoal, you can use the combinator @{ML_ind EVERY' in
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   975
  Tactical}. For example the function @{ML foo_tac'} from
3d27d77c351f added RANGE
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
   976
  page~\pageref{tac:footacprime} can also be written as:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   977
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   978
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   979
ML %grayML{*fun foo_tac'' ctxt = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   980
  EVERY' [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}], 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   981
          assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]*}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
   982
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   983
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   984
  There is even another way of implementing this tactic: in automatic proof
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   985
  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
   986
  are often long lists of tactics that are applied to the first
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   987
  subgoal. Instead of writing the code above and then calling @{ML "foo_tac'' @{context} 1"}, 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   988
  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
   989
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   990
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   991
ML %grayML{*fun foo_tac1 ctxt = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   992
  EVERY1 [eresolve_tac ctxt [@{thm disjE}], resolve_tac ctxt [@{thm disjI2}], 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   993
          assume_tac ctxt, resolve_tac ctxt [@{thm disjI1}], assume_tac ctxt]*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   994
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   995
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   996
  and call @{ML foo_tac1}.  
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   997
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
   998
  With the combinators @{ML THEN'}, @{ML EVERY'} and @{ML_ind  EVERY1 in Tactical} it must be
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   999
  guaranteed that all component tactics successfully apply; otherwise the
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1000
  whole tactic will fail. If you rather want to try out a number of tactics,
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  1001
  then you can use the combinator @{ML_ind  ORELSE' in Tactical} for two tactics, and @{ML_ind
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  1002
   FIRST' in Tactical} (or @{ML_ind  FIRST1 in Tactical}) for a list of tactics. For example, the tactic
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1003
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1004
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1005
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1006
ML %grayML{*fun orelse_xmp_tac ctxt = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1007
  resolve_tac ctxt [@{thm disjI1}] ORELSE' resolve_tac ctxt [@{thm conjI}]*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
  1008
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1009
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1010
  will first try out whether theorem @{text disjI} applies and in case of failure 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1011
  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
  1012
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1013
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1014
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1015
  shows "True \<and> False" and "Foo \<or> Bar"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1016
apply(tactic {* orelse_xmp_tac @{context} 2 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1017
apply(tactic {* orelse_xmp_tac @{context} 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1018
txt {* which results in the goal state
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1019
       \begin{isabelle}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1020
       @{subgoals [display]} 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1021
       \end{isabelle}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1022
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1024
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1027
  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
  1028
  as follows:
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1029
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1030
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1031
ML %grayML{*fun select_tac' ctxt = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1032
  FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}], 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1033
          resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{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
  1034
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1035
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
  1036
  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
  1037
  we must include @{ML all_tac} at the end of the list, otherwise the tactic will
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1038
  fail if no theorem applies (we also have to wrap @{ML all_tac} using the 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1039
  @{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
  1040
  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
  1041
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1042
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1043
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1044
  shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1045
apply(tactic {* select_tac' @{context} 4 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1046
apply(tactic {* select_tac' @{context} 3 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1047
apply(tactic {* select_tac' @{context} 2 *})
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1048
apply(tactic {* select_tac' @{context} 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1049
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1050
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1051
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1052
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1053
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1054
text {* 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1055
  Since such repeated applications of a tactic to the reverse order of 
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1056
  \emph{all} subgoals is quite common, there is the tactic combinator 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  1057
  @{ML_ind  ALLGOALS in Tactical} 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
  1058
  write: *}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1059
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1060
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1061
  shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow>C" and "\<forall>x. D x" and "E \<Longrightarrow> F"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1062
apply(tactic {* ALLGOALS (select_tac' @{context}) *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1063
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1064
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1065
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1066
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1067
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1068
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1069
  Remember that we chose to implement @{ML select_tac'} so that it 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1070
  always  succeeds by fact of having @{ML all_tac} at the end of the tactic
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  1071
  list. The same can be achieved with the tactic combinator @{ML_ind  TRY in Tactical}.
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1072
  For example:
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1073
*}
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1074
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1075
ML %grayML{*fun select_tac'' ctxt = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1076
 TRY o FIRST' [resolve_tac ctxt [@{thm conjI}], resolve_tac ctxt [@{thm impI}], 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1077
               resolve_tac ctxt [@{thm notI}], resolve_tac ctxt [@{thm allI}]]*}
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1078
text_raw{*\label{tac:selectprimeprime}*}
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1079
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1080
text {*
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1081
  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
  1082
  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
  1083
  unchanged. This, however, can be potentially very confusing when visible to 
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1084
  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
  1085
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1086
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1087
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1088
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1089
  shows "E \<Longrightarrow> F"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1090
apply(tactic {* select_tac' @{context} 1 *})
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1091
txt{* \begin{minipage}{\textwidth}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1092
      @{subgoals [display]}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1093
      \end{minipage} *}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1094
(*<*)oops(*>*)
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1095
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1096
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1097
  In this case no theorem applies. But because we wrapped the tactic in a @{ML
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1098
  TRY}, it does not fail anymore. The problem is that for the user there is
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1099
  little chance to see whether progress in the proof has been made, or not. By
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1100
  convention therefore, tactics visible to the user should either change
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1101
  something or fail.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1102
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1103
  To comply with this convention, we could simply delete the @{ML "K all_tac"}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1104
  in @{ML select_tac'} or delete @{ML TRY} from @{ML select_tac''}. But for
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1105
  the sake of argument, let us suppose that this deletion is \emph{not} an
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1106
  option. In such cases, you can use the combinator @{ML_ind CHANGED in
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1107
  Tactical} to make sure the subgoal has been changed by the tactic. Because
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1108
  now
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1109
*}
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1110
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1111
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1112
  shows "E \<Longrightarrow> F"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1113
apply(tactic {* CHANGED (select_tac' @{context} 1) *})(*<*)?(*>*)
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1114
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
  1115
  \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
  1116
  @{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
  1117
  @{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
  1118
  \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
  1119
*}(*<*)oops(*>*)
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1120
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1121
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1122
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1123
  We can further extend the @{ML select_tac}s so that they not just apply to the topmost
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1124
  connective, but also to the ones immediately ``underneath'', i.e.~analyse the goal 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1125
  completely. For this you can use the tactic combinator @{ML_ind REPEAT in Tactical}. As an example 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1126
  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
  1127
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1128
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1129
ML %grayML{*fun repeat_xmp_tac ctxt = REPEAT (CHANGED (select_tac' ctxt 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
  1130
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1131
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
  1132
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1133
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1134
  shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1135
apply(tactic {* repeat_xmp_tac @{context} *})
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1136
txt{* produces
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1137
      \begin{isabelle}
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
  1138
      @{subgoals [display]}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1139
      \end{isabelle} *}
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
  1140
(*<*)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
  1141
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1142
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1143
  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
  1144
  because otherwise @{ML REPEAT} runs into an infinite loop (it applies the
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1145
  tactic as long as it succeeds). The function @{ML_ind REPEAT1 in Tactical}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1146
  is similar, but runs the tactic at least once (failing if this is not
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1147
  possible).
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
  1148
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1149
  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
  1150
  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
  1151
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1152
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1153
ML %grayML{*fun repeat_xmp_tac' ctxt = REPEAT o CHANGED o select_tac' ctxt*}
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
  1154
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1155
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
  1156
  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
  1157
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1158
  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
  1159
  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
  1160
  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
  1161
  is applied repeatedly only to the first subgoal. To analyse also all
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  1162
  resulting subgoals, you can use the tactic combinator @{ML_ind  REPEAT_ALL_NEW in Tactical}. 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1163
  Supposing the 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
  1164
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1165
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1166
ML %grayML{*fun repeat_all_new_xmp_tac ctxt = REPEAT_ALL_NEW (CHANGED o select_tac' ctxt)*}
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
  1167
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1168
text {* 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1169
  you can 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
  1170
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1171
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1172
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1173
  shows "((\<not>A) \<and> (\<forall>x. B x)) \<and> (C \<longrightarrow> D)"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1174
apply(tactic {* repeat_all_new_xmp_tac @{context} 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
  1175
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
  1176
      @{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
  1177
      \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
  1178
(*<*)oops(*>*)
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1179
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
  1180
text {*
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1181
  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
  1182
  include in @{ML select_tac'}. 
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1183
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1184
  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
  1185
  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
  1186
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1187
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1188
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1189
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1190
  shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1191
apply(tactic {* eresolve_tac @{context} [@{thm disjE}] 1 *})
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1192
txt{* applies the rule to the first assumption yielding the goal state:
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1193
      \begin{isabelle}
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
  1194
      @{subgoals [display]}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1195
      \end{isabelle}
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1196
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1197
      After typing
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1198
  *}
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
  1199
(*<*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1200
oops
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1201
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1202
  shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1203
apply(tactic {* eresolve_tac @{context} [@{thm disjE}] 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
  1204
(*>*)
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1205
back
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1206
txt{* the rule now applies to the second assumption.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1207
      \begin{isabelle}
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
  1208
      @{subgoals [display]}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1209
      \end{isabelle} *}
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
  1210
(*<*)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
  1211
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1212
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
  1213
  Sometimes this leads to confusing behaviour of tactics and also has
109
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1214
  the potential to explode the search space for tactics.
b4234e8a0202 polished
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
  1215
  These problems can be avoided by prefixing the tactic with the tactic 
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  1216
  combinator @{ML_ind  DETERM in Tactical}.
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
  1217
*}
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1218
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1219
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1220
  shows "\<lbrakk>P1 \<or> Q1; P2 \<or> Q2\<rbrakk> \<Longrightarrow> R"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1221
apply(tactic {* DETERM (eresolve_tac @{context} [@{thm disjE}] 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
  1222
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
  1223
      @{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
  1224
      \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
  1225
(*<*)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
  1226
text {*
118
5f003fdf2653 polished and added more material to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
  1227
  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
  1228
  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
  1229
  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
  1230
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 107
diff changeset
  1231
  \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
  1232
  @{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
  1233
  @{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
  1234
  \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
  1235
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1236
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  1237
  \footnote{\bf FIXME: say something about @{ML_ind COND in Tactical} and COND'}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1238
  \footnote{\bf FIXME: PARALLEL-CHOICE PARALLEL-GOALS}
238
29787dcf7b2e added something about TRY and TRYALL
Christian Urban <urbanc@in.tum.de>
parents: 232
diff changeset
  1239
  
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
  1240
  \begin{readmore}
289
08ffafe2585d adapted to changes in Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 288
diff changeset
  1241
  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
  1242
  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
  1243
  \end{readmore}
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1244
*}
107
258ce361ba1b polished and more material in the tactic chapter
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
  1245
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1246
text {*
313
1ca2f41770cc polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
  1247
  \begin{exercise}\label{ex:dyckhoff}
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1248
  Dyckhoff presents in \cite{Dyckhoff92} inference rules of a sequent
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1249
  calculus, called G4ip, for intuitionistic propositional logic. The
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1250
  interesting feature of this calculus is that no contraction rule is needed
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1251
  in order to be complete. As a result the rules can be applied exhaustively, which
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1252
  in turn leads to simple decision procedure for propositional intuitionistic logic. 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1253
  The task is to implement this decision procedure as a tactic. His rules are
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1254
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1255
  \begin{center}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1256
  \def\arraystretch{2.3}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1257
  \begin{tabular}{cc}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1258
  \infer[Ax]{A,\varGamma \Rightarrow A}{} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1259
  \infer[False]{False,\varGamma \Rightarrow G}{}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1260
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1261
  \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
  1262
  \infer[\wedge_R]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1263
  {\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
  1264
  
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1265
  \infer[\vee_L]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1266
  {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
  1267
  \infer[\vee_{R_1}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1268
  {\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow A} \hspace{3mm}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1269
  \infer[\vee_{R_2}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1270
  {\varGamma \Rightarrow A \vee B}{\varGamma \Rightarrow B}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1271
  
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1272
  \infer[\longrightarrow_{L_1}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1273
  {A\longrightarrow B, A, \varGamma \Rightarrow G}{B, A, \varGamma \Rightarrow G} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1274
  \infer[\longrightarrow_R]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1275
  {\varGamma \Rightarrow A\longrightarrow B}{A,\varGamma \Rightarrow B}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1276
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1277
  \infer[\longrightarrow_{L_2}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1278
  {(C \wedge D)\longrightarrow B, \varGamma \Rightarrow G}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1279
  {C\longrightarrow (D \longrightarrow B), \varGamma \Rightarrow G} &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1280
  
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1281
  \infer[\longrightarrow_{L_3}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1282
  {(C \vee D)\longrightarrow B, \varGamma \Rightarrow G}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1283
  {C\longrightarrow B, D\longrightarrow B, \varGamma \Rightarrow G}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1284
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1285
  \multicolumn{2}{c}{
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1286
  \infer[\longrightarrow_{L_4}]
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1287
  {(C \longrightarrow D)\longrightarrow B, \varGamma \Rightarrow G}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1288
  {D\longrightarrow B, \varGamma \Rightarrow C \longrightarrow D &
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1289
  B, \varGamma \Rightarrow G}}\\
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1290
  \end{tabular}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1291
  \end{center}
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
  1292
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1293
  Note that in Isabelle right rules need to be implemented as
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1294
  introduction rule, the left rules as elimination rules. You have to to
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1295
  prove separate theorems corresponding to $\longrightarrow_{L_{1..4}}$. The
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1296
  tactic should explore all possibilites of applying these rules to a
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1297
  propositional formula until a goal state is reached in which all subgoals
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1298
  are discharged. For this you can use the tactic combinator @{ML_ind
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1299
  DEPTH_SOLVE in Search} in the structure @{ML_struct Search}.
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1300
  \end{exercise}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1301
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1302
  \begin{exercise}
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1303
  Add to the sequent calculus from the previous exercise also rules for 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1304
  equality and run your tactic on  the de Bruijn formulae discussed 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1305
  in Exercise~\ref{ex:debruijn}.
313
1ca2f41770cc polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
  1306
  \end{exercise}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1307
*}
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1308
388
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 384
diff changeset
  1309
section {* Simplifier Tactics\label{sec:simplifier} *}
105
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1310
f49dc7e96235 added more to the Tactical section
Christian Urban <urbanc@in.tum.de>
parents: 104
diff changeset
  1311
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1312
  A lot of convenience in reasoning with Isabelle derives from its
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1313
  powerful simplifier. We will describe it in this section. However,
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1314
  due to its complexity, we can mostly only scratch the surface. 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1315
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1316
  The power of the simplifier is a strength and a weakness at the same time,
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1317
  because you can easily make the simplifier run into a loop and in general
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1318
  its behaviour can be difficult to predict. There is also a multitude of
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1319
  options that you can configure to change the behaviour of the
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1320
  simplifier. There are the following five main tactics behind the simplifier
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1321
  (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
  1322
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1323
  \begin{isabelle}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1324
  \begin{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1325
  \begin{tabular}{l@ {\hspace{2cm}}l}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1326
   @{ML_ind simp_tac in Simplifier}            & @{text "(simp (no_asm))"} \\
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1327
   @{ML_ind asm_simp_tac in Simplifier}        & @{text "(simp (no_asm_simp))"} \\
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1328
   @{ML_ind full_simp_tac in Simplifier}       & @{text "(simp (no_asm_use))"} \\
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1329
   @{ML_ind asm_lr_simp_tac in Simplifier}     & @{text "(simp (asm_lr))"} \\
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1330
   @{ML_ind asm_full_simp_tac in Simplifier}   & @{text "(simp)"}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1331
  \end{tabular}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1332
  \end{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1333
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1334
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1335
  All these 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
  1336
  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
  1337
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1338
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1339
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1340
  shows "Suc (1 + 2) < 3 + 2"
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1341
apply(simp)
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1342
done
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1343
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1344
text {*
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1345
  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
  1346
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1347
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1348
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1349
  shows "Suc (1 + 2) < 3 + 2"
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1350
apply(tactic {* asm_full_simp_tac @{context} 1 *})
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1351
done
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1352
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1353
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1354
  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
  1355
  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
  1356
  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
  1357
  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
  1358
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1359
  \footnote{\bf FIXME: show rewriting of cterms}
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
  1360
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  1361
  There is one restriction you have to keep in mind when using the simplifier:
413
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
  1362
  it can only deal with rewriting rules whose left-hand sides are higher order 
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
  1363
  pattern (see Section \ref{sec:univ} on unification). Whether a term is a pattern
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
  1364
  or not can be tested with the function @{ML_ind pattern in Pattern} from
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
  1365
  the structure @{ML_struct Pattern}. If a rule is not a pattern and you want
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
  1366
  to use it for rewriting, then you have to use simprocs or conversions, both 
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
  1367
  of which we shall describe in the next section.
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  1368
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1369
  When using the simplifier, the crucial information you have to provide is
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1370
  the simpset. If this information is not handled with care, then, as
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1371
  mentioned above, the simplifier can easily run into a loop. Therefore a good
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1372
  rule of thumb is to use simpsets that are as minimal as possible. It might
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1373
  be surprising that a simpset is more complex than just a simple collection
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1374
  of theorems. One reason for the complexity is that the simplifier must be
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1375
  able to rewrite inside terms and should also be able to rewrite according to
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1376
  theorems that have premises.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1377
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1378
  The rewriting inside terms requires congruence theorems, which
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1379
  are typically meta-equalities of the form
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1380
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1381
  \begin{isabelle}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1382
  \begin{center}
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1383
  \mbox{\inferrule{@{text "t\<^sub>1 \<equiv> s\<^sub>1 \<dots> t\<^sub>n \<equiv> s\<^sub>n"}}
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1384
                  {@{text "constr t\<^sub>1\<dots>t\<^sub>n \<equiv> constr s\<^sub>1\<dots>s\<^sub>n"}}}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1385
  \end{center}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1386
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1387
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1388
  with @{text "constr"} being a constant, like @{const "If"}, @{const "Let"}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1389
  and so on. Every simpset contains only one congruence rule for each
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1390
  term-constructor, which however can be overwritten. The user can declare
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1391
  lemmas to be congruence rules using the attribute @{text "[cong]"}. Note that 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1392
  in HOL these congruence theorems are usually stated as equations, which are 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1393
  then internally transformed into meta-equations.
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1394
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1395
  The rewriting with theorems involving premises requires what is in Isabelle
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1396
  called a subgoaler, a solver and a looper. These can be arbitrary tactics
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1397
  that can be installed in a simpset and which are executed at various stages
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1398
  during simplification. 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1399
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1400
  Simpsets can also include simprocs, which produce rewrite rules on
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1401
  demand according to a pattern (see next section for a detailed description
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1402
  of simpsets). Another component are split-rules, which can simplify for
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1403
  example the ``then'' and ``else'' branches of if-statements under the
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1404
  corresponding preconditions.
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1405
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1406
  \begin{readmore}
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1407
  For more information about the simplifier see @{ML_file "Pure/raw_simplifier.ML"}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1408
  and @{ML_file "Pure/simplifier.ML"}. The generic splitter is implemented in 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1409
  @{ML_file "Provers/splitter.ML"}.
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1410
  \end{readmore}
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1411
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1412
  
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1413
  \footnote{\bf FIXME: Find the right place to mention this: Discrimination 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1414
  nets are implemented in @{ML_file "Pure/net.ML"}.}
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 158
diff changeset
  1415
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1416
  The most common combinators for modifying simpsets are:
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1417
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1418
  \begin{isabelle}
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1419
  \begin{tabular}{l@ {\hspace{10mm}}l}
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1420
  @{ML_ind addsimps in Raw_Simplifier}    & @{ML_ind delsimps in Raw_Simplifier}\\
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1421
  @{ML_ind addsimprocs in Raw_Simplifier} & @{ML_ind delsimprocs in Raw_Simplifier}\\
503
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 489
diff changeset
  1422
  @{ML_ind add_cong in Raw_Simplifier}    & @{ML_ind del_cong in Raw_Simplifier}\\
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1423
  \end{tabular}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1424
  \end{isabelle}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1425
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1426
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1427
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1428
text_raw {*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1429
\begin{figure}[t]
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1430
\begin{minipage}{\textwidth}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1431
\begin{isabelle}*}
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
  1432
ML %grayML{*fun print_ss ctxt ss =
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1433
let
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1434
  val {simps, congs, procs, ...} = Raw_Simplifier.dest_ss ss
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1435
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1436
  fun name_sthm (nm, thm) =
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1437
    Pretty.enclose (nm ^ ": ") "" [pretty_thm_no_vars ctxt thm]
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1438
  fun name_cthm ((_, nm), thm) =
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
  1439
    Pretty.enclose (nm ^ ": ") "" [pretty_thm_no_vars ctxt thm]
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1440
  fun name_trm (nm, trm) =
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1441
    Pretty.enclose (nm ^ ": ") "" [pretty_terms ctxt trm]
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
  1442
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1443
  val pps = [Pretty.big_list "Simplification rules:" (map name_sthm simps),
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1444
             Pretty.big_list "Congruences rules:" (map name_cthm congs),
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1445
             Pretty.big_list "Simproc patterns:" (map name_trm procs)]
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1446
in
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
  1447
  pps |> Pretty.chunks
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 437
diff changeset
  1448
      |> pwriteln
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1449
end*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1450
text_raw {* 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1451
\end{isabelle}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  1452
\end{minipage}
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1453
\caption{The function @{ML_ind dest_ss in Raw_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
  1454
  all printable information stored in a simpset. We are here only interested in the 
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1455
  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
  1456
\end{figure} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1457
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
  1458
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1459
text {* 
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1460
  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
  1461
  prints out some parts of a simpset. If you use it to print out the components of the
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1462
  empty simpset, i.e., @{ML_ind empty_ss in Raw_Simplifier}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1463
  
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1464
  @{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
  1465
  "print_ss @{context} empty_ss"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1466
"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
  1467
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
  1468
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1469
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1470
  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
  1471
  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
  1472
  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
  1473
*}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1474
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1475
ML %grayML{*val ss1 = put_simpset empty_ss @{context} addsimps [@{thm Diff_Int} RS @{thm eq_reflection}] *}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1476
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1477
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1478
  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
  1479
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1480
  @{ML_response_fake [display,gray]
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1481
  "print_ss @{context} (Raw_Simplifier.simpset_of ss1)"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1482
"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
  1483
  ??.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
  1484
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
  1485
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1486
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1487
  \footnote{\bf FIXME: Why does it print out ??.unknown}
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
  1488
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 552
diff changeset
  1489
  Adding also the congruence rule @{thm [source] strong_INF_cong} 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1490
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1491
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 552
diff changeset
  1492
ML %grayML{*val ss2 = ss1 |> Simplifier.add_cong (@{thm strong_INF_cong} RS @{thm eq_reflection}) *}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1493
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1494
text {*
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1495
  gives
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1496
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1497
  @{ML_response_fake [display,gray]
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1498
  "print_ss @{context} (Raw_Simplifier.simpset_of ss2)"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1499
"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
  1500
  ??.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
  1501
Congruences rules:
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 552
diff changeset
  1502
  Complete_Lattices.Inf_class.INFIMUM: 
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 552
diff changeset
  1503
    \<lbrakk>A1 = B1; \<And>x. x \<in> B1 =simp=> C1 x = D1 x\<rbrakk> \<Longrightarrow> INFIMUM A1 C1 \<equiv> INFIMUM B1 D1
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
  1504
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1505
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1506
  Notice that we had to add these lemmas as meta-equations. The @{ML empty_ss} 
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1507
  expects this form of the simplification and congruence rules. This is
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1508
  different, if we use for example the simpset @{ML HOL_basic_ss} (see below), 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1509
  where rules are usually added as equation. However, even 
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1510
  when adding these lemmas to @{ML empty_ss} we do not end up with anything useful yet.
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1511
  In the context of HOL, the first really useful simpset is @{ML_ind
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1512
  HOL_basic_ss in Simpdata}. While printing out the components of this simpset
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1513
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1514
  @{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
  1515
  "print_ss @{context} HOL_basic_ss"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1516
"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
  1517
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
  1518
Simproc patterns:"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1519
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1520
  also produces ``nothing'', the printout is somewhat misleading. In fact
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1521
  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
  1522
  form  
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1523
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1524
  \begin{isabelle}
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1525
  @{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
  1526
  \end{isabelle}
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1527
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1528
  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
  1529
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1530
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1531
lemma 
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1532
 shows "True" 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1533
  and  "t = t" 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1534
  and  "t \<equiv> t" 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1535
  and  "False \<Longrightarrow> Foo" 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1536
  and  "\<lbrakk>A; B; C\<rbrakk> \<Longrightarrow> A"
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1537
apply(tactic {* ALLGOALS (simp_tac (put_simpset HOL_basic_ss @{context})) *})
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1538
done
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1539
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1540
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1541
  This behaviour is not because of simplification rules, but how the subgoaler, solver 
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  1542
  and looper are set up in @{ML HOL_basic_ss}.
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1543
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1544
  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
  1545
  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
  1546
  connectives in HOL. 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1547
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1548
  @{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
  1549
  "print_ss @{context} HOL_ss"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1550
"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
  1551
  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
  1552
  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
  1553
  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
  1554
  \<dots>
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1555
Congruences rules:
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1556
  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
  1557
    \<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
  1558
  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
  1559
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
  1560
  \<dots>"}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1561
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1562
  \begin{readmore}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1563
  The simplifier for HOL is set up in @{ML_file "HOL/Tools/simpdata.ML"}.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1564
  The simpset @{ML HOL_ss} is implemented in @{ML_file "HOL/HOL.thy"}.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1565
  \end{readmore}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1566
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1567
  The simplifier is often used to unfold definitions in a proof. For this the
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1568
  simplifier implements the tactic @{ML_ind rewrite_goal_tac in Raw_Simplifier}.\footnote{\bf FIXME: 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1569
  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
  1570
  definition
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1571
*}
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1572
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1573
definition "MyTrue \<equiv> True"
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1574
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1575
text {*
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1576
  then we can use this tactic to unfold the definition of this constant.
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1577
*}
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  1578
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1579
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1580
  shows "MyTrue \<Longrightarrow> True"
552
82c482467d75 updated to latest isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 551
diff changeset
  1581
apply(tactic {* rewrite_goal_tac @{context} @{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
  1582
txt{* producing the goal state
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1583
      \begin{isabelle}
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1584
      @{subgoals [display]}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1585
      \end{isabelle} *}
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1586
(*<*)oops(*>*)
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1587
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1588
text {*
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1589
  If you want to unfold definitions in \emph{all} subgoals, not just one, 
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1590
  then use the the tactic @{ML_ind rewrite_goals_tac in Raw_Simplifier}.
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1591
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1592
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1593
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1594
text_raw {*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1595
\begin{figure}[p]
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1596
\begin{boxedminipage}{\textwidth}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1597
\begin{isabelle} *}
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
  1598
type_synonym  prm = "(nat \<times> nat) list"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1599
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
  1600
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1601
overloading 
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1602
  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
  1603
  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
  1604
  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
  1605
begin
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1606
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1607
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
  1608
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1609
 "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
  1610
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1611
primrec perm_nat
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1612
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1613
  "perm_nat [] c = c"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1614
| "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
  1615
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1616
fun perm_prod 
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1617
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1618
 "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
  1619
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1620
primrec perm_list
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1621
where
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1622
  "perm_list pi [] = []"
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1623
| "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
  1624
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1625
end
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1626
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1627
lemma perm_append[simp]:
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1628
  fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1629
  shows "((pi\<^sub>1@pi\<^sub>2)\<bullet>c) = (pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c))"
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1630
by (induct pi\<^sub>1) (auto) 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1631
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1632
lemma perm_bij[simp]:
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1633
  fixes c d::"nat" and pi::"prm"
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1634
  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
  1635
by (induct pi) (auto)
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1636
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1637
lemma perm_rev[simp]:
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1638
  fixes c::"nat" and pi::"prm"
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1639
  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
  1640
by (induct pi arbitrary: c) (auto)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1641
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1642
lemma perm_compose:
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1643
fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1644
shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c) = (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>(pi\<^sub>1\<bullet>c)" 
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1645
by (induct pi\<^sub>2) (auto)
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1646
text_raw {*
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1647
\end{isabelle}
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 172
diff changeset
  1648
\end{boxedminipage}
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1649
\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
  1650
  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
  1651
  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
  1652
  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
  1653
  \label{fig:perms}}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1654
\end{figure} *}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1655
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1656
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1657
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1658
  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
  1659
  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
  1660
  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
  1661
  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
  1662
  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
  1663
  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
  1664
  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
  1665
  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
  1666
  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
  1667
  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
  1668
  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
  1669
  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
  1670
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1671
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1672
lemma 
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1673
fixes c d::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1674
  shows "pi\<^sub>1\<bullet>(c, pi\<^sub>2\<bullet>((rev pi\<^sub>1)\<bullet>d)) = (pi\<^sub>1\<bullet>c, (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>d)"
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1675
apply(simp)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1676
apply(rule trans)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1677
apply(rule perm_compose)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1678
apply(simp)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1679
done 
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1680
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1681
text {*
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1682
  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
  1683
  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
  1684
  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
  1685
  assume the auxiliary constant is
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1686
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1687
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1688
definition
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1689
  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
  1690
where
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1691
  "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
  1692
162
3fb9f820a294 some additions to the simplifier section and general tuning
Christian Urban <urbanc@in.tum.de>
parents: 161
diff changeset
  1693
text {* Now the two lemmas *}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1694
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1695
lemma perm_aux_expand:
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1696
  fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1697
  shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>c) = pi\<^sub>1 \<bullet>aux (pi\<^sub>2\<bullet>c)" 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1698
unfolding perm_aux_def by (rule refl)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1699
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1700
lemma perm_compose_aux:
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1701
  fixes c::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1702
  shows "pi\<^sub>1\<bullet>(pi\<^sub>2\<bullet>aux c) = (pi\<^sub>1\<bullet>pi\<^sub>2) \<bullet>aux (pi\<^sub>1\<bullet>c)" 
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1703
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
  1704
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1705
text {* 
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1706
  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
  1707
  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
  1708
  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
  1709
  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
  1710
  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
  1711
  into three phases without the user noticing anything about the auxiliary 
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1712
  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
  1713
  lemma @{thm [source] "perm_aux_expand"} (Line 9);
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1714
  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
  1715
  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
  1716
  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
  1717
  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
  1718
*}
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1719
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1720
ML %linenosgray{*fun perm_simp_tac ctxt =
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1721
let
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1722
  val thms1 = [@{thm perm_aux_expand}]
229
abc7f90188af permutation example uses now recent infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1723
  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
  1724
               @{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
  1725
               @{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
  1726
  val thms3 = [@{thm perm_aux_def}]
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1727
  val ss = put_simpset HOL_basic_ss ctxt
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1728
in
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1729
  simp_tac (ss addsimps thms1)
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1730
  THEN' simp_tac (ss addsimps thms2)
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1731
  THEN' simp_tac (ss addsimps thms3)
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1732
end*}
153
c22b507e1407 general polishing; added versioning to the document
Christian Urban <urbanc@in.tum.de>
parents: 152
diff changeset
  1733
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1734
text {*
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1735
  For all three phases we have to build simpsets adding specific lemmas. As is sufficient
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1736
  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
  1737
  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
  1738
*}
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1739
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1740
lemma 
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1741
  fixes c d::"nat" and pi\<^sub>1 pi\<^sub>2::"prm"
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  1742
  shows "pi\<^sub>1\<bullet>(c, pi\<^sub>2\<bullet>((rev pi\<^sub>1)\<bullet>d)) = (pi\<^sub>1\<bullet>c, (pi\<^sub>1\<bullet>pi\<^sub>2)\<bullet>d)"
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1743
apply(tactic {* perm_simp_tac @{context} 1 *})
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1744
done
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1745
152
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  1746
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1747
text {*
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
  1748
  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
  1749
  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
  1750
  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
  1751
  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
  1752
  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
  1753
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1754
  (FIXME: Is it interesting to say something about @{term "(=simp=>)"}?)
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1755
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1756
  (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
  1757
  do with weak congruence constants?)
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
  1758
240
d111f5988e49 replaced explode by Symbol.explode
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
  1759
  (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
  1760
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
  1761
  (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
  1762
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1763
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1764
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1765
section {* Simprocs *}
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
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1768
  In Isabelle you can also implement custom simplification procedures, called
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1769
  \emph{simprocs}. Simprocs can be triggered by the simplifier on a specified
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1770
  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
  1771
  cases where a rewriting rule must be produced on ``demand'' or when
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1772
  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
  1773
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1774
  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
  1775
  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
  1776
  you can use the function:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1777
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1778
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1779
ML %linenosgray{*fun fail_simproc ctxt redex = 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1780
let
543
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 541
diff changeset
  1781
  val _ = pwriteln (Pretty.block 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 541
diff changeset
  1782
    [Pretty.str "The redex: ", pretty_cterm ctxt redex])
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1783
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1784
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1785
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1786
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1787
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1788
  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
  1789
  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
  1790
  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
  1791
  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
  1792
  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
  1793
  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
  1794
  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
  1795
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1796
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1797
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
  1798
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1799
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1800
  where the second argument specifies the pattern and the right-hand side
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1801
  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
  1802
  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
  1803
  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
  1804
  it fires on the lemma:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1805
*}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1806
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1807
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1808
  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
  1809
apply(simp)
378
8d160d79b48c section about matching and unification of types
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
  1810
txt{*
8d160d79b48c section about matching and unification of types
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
  1811
  \begin{minipage}{\textwidth}
8d160d79b48c section about matching and unification of types
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
  1812
  \small@{text "> The redex: Suc 0"}\\
213
e60dbcba719d some polishing
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
  1813
  @{text "> The redex: Suc 0"}\\
378
8d160d79b48c section about matching and unification of types
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
  1814
  \end{minipage}*}(*<*)oops(*>*)
8d160d79b48c section about matching and unification of types
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
  1815
8d160d79b48c section about matching and unification of types
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
  1816
text {* 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1817
  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
  1818
  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
  1819
  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
  1820
  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
  1821
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1822
  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
  1823
  \isacommand{declare}-statement. For example the simproc will be deleted
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1824
  with the declaration
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1825
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1826
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1827
declare [[simproc del: fail]]
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1828
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1829
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1830
  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
  1831
  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
  1832
  as follows:
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1833
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1834
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1835
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1836
  shows "Suc 0 = 1"
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1837
apply(tactic {* simp_tac (put_simpset 
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1838
  HOL_basic_ss @{context} addsimprocs [@{simproc fail}]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1839
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1840
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1841
text {*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1842
  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
  1843
  left unchanged. 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1844
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
  1845
  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
  1846
  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
  1847
  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
  1848
  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
  1849
  to
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1850
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1851
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1852
ML %grayML{*fun fail_simproc' _ ctxt redex = 
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1853
let
543
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 541
diff changeset
  1854
  val _ = pwriteln (Pretty.block 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1855
    [Pretty.str "The redex:", pretty_cterm ctxt redex])
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1856
in
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1857
  NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1858
end*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1859
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1860
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1861
  Here the redex is given as a @{ML_type term}, instead of a @{ML_type cterm}
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
  1862
  (therefore we printing it out using the function @{ML pretty_term in Pretty}).
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  1863
  We can turn this function into a proper simproc using the function 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1864
  @{ML Simplifier.make_simproc}:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1865
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1866
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1867
ML %grayML{*val fail' = 
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1868
  Simplifier.make_simproc @{context} "fail_simproc'"
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1869
    {lhss = [@{term "Suc n"}], proc = fail_simproc'}*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1870
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1871
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1872
  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
  1873
  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
  1874
  Now the simproc is set up and can be explicitly added using
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  1875
  @{ML_ind addsimprocs in Raw_Simplifier} to a simpset whenever
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1876
  needed. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1877
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1878
  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
  1879
  see this in the proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1880
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1881
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1882
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1883
  shows "Suc (Suc 0) = (Suc 1)"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1884
apply(tactic {* simp_tac ((put_simpset HOL_basic_ss @{context}) addsimprocs [fail']) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1885
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1886
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1887
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  1888
  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
  1889
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1890
@{text [display]
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1891
"> Suc 0
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1892
> Suc (Suc 0) 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1893
> Suc 1"}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1894
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1895
  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
  1896
  rewrites terms according to the equation:
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1897
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1898
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1899
lemma plus_one: 
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1900
  shows "Suc n \<equiv> n + 1" by simp
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1901
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1902
text {*
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1903
  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
  1904
  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
  1905
  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
  1906
  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
  1907
  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
  1908
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1909
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1910
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1911
ML %grayML{*fun plus_one_simproc _ ctxt redex =
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1912
  case Thm.term_of redex of
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1913
    @{term "Suc 0"} => NONE
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1914
  | _ => SOME @{thm plus_one}*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1915
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1916
text {*
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1917
  and set up the simproc as follows.
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1918
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1919
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1920
ML %grayML{*val plus_one =
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1921
  Simplifier.make_simproc @{context} "sproc +1"
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1922
    {lhss = [@{term "Suc n"}], proc = plus_one_simproc}*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1923
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1924
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1925
  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
  1926
  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
  1927
  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
  1928
  in the following proof
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1929
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1930
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1931
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  1932
  shows "P (Suc (Suc (Suc 0))) (Suc 0)"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1933
apply(tactic {* simp_tac (put_simpset HOL_basic_ss @{context} addsimprocs [plus_one]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1934
txt{*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1935
  where the simproc produces the goal state
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1936
  \begin{isabelle}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1937
  @{subgoals[display]}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  1938
  \end{isabelle}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1939
*}  
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1940
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1941
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1942
text {*
133
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
  1943
  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
  1944
  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
  1945
  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
  1946
  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
  1947
  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
  1948
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1949
  Next let us implement a simproc that replaces terms of the form @{term "Suc n"}
232
0d03e1304ef6 minor changes
griff
parents: 231
diff changeset
  1950
  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
  1951
  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
  1952
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1953
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
  1954
ML %grayML{*fun dest_suc_trm ((Const (@{const_name "Suc"}, _)) $ t) = 1 + dest_suc_trm t
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  1955
  | 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
  1956
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1957
text {* 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  1958
  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
  1959
  (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
  1960
  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
  1961
  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
  1962
  @{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
  1963
*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1964
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1965
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
  1966
let
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1967
  val num = HOLogic.mk_number @{typ "nat"} n
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1968
  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
  1969
in
214
7e04dc2368b0 updated to latest Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 213
diff changeset
  1970
  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
  1971
end*}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1972
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1973
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1974
  From the integer value it generates the corresponding number term, called 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1975
  @{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
  1976
  (Line 4), which it proves by the arithmetic tactic in Line 6. 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1977
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 217
diff changeset
  1978
  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
  1979
  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
  1980
  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
  1981
  want to prove here, the simpset @{text "num_ss"} should suffice.
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1982
*}
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 130
diff changeset
  1983
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
  1984
ML %grayML{*fun get_thm_alt ctxt (t, n) =
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1985
let
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1986
  val num = HOLogic.mk_number @{typ "nat"} n
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1987
  val goal = Logic.mk_equals (t, num)
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  1988
  val num_ss = put_simpset HOL_ss ctxt addsimps @{thms semiring_norm}
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1989
in
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1990
  Goal.prove ctxt [] [] goal (K (simp_tac num_ss 1))
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1991
end*}
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  1992
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1993
text {*
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1994
  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
  1995
  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
  1996
  what happens with @{ML arith_tac in Arith_Data}, especially in more complicated 
231
f4c9dd7bcb28 ran spell-checker
griff
parents: 230
diff changeset
  1997
  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
  1998
  that is sufficiently powerful to solve every instance of the lemmas
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  1999
  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
  2000
  ``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
  2001
  way than tracing the simplifier to obtain the lemmas that are successfully 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  2002
  applied during simplification. Alas, there is none.} 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  2003
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  2004
  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
  2005
  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
  2006
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2007
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2008
ML %grayML{*fun nat_number_simproc _ ctxt ct =
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2009
  SOME (get_thm_alt ctxt (Thm.term_of ct, dest_suc_trm (Thm.term_of ct)))
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 543
diff changeset
  2010
  handle TERM _ => NONE *}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2011
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2012
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2013
  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
  2014
  @{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
  2015
  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
  2016
  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
  2017
*}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2018
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2019
ML %grayML{*val nat_number =
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2020
  Simplifier.make_simproc @{context} "nat_number"
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2021
    {lhss = [@{term "Suc n"}], proc = nat_number_simproc}*}
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2022
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2023
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2024
text {* 
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2025
  Now in the lemma
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2026
  *}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2027
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  2028
lemma 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  2029
  shows "P (Suc (Suc 2)) (Suc 99) (0::nat) (Suc 4 + Suc 0) (Suc (0 + 0))"
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2030
apply(tactic {* simp_tac (put_simpset HOL_ss @{context} addsimprocs [nat_number]) 1*})
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2031
txt {* 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2032
  you obtain the more legible goal state
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2033
  \begin{isabelle}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2034
  @{subgoals [display]}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2035
  \end{isabelle}*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2036
(*<*)oops(*>*)
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2037
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2038
text {*
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  2039
  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
  2040
  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
  2041
  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
  2042
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2043
  \begin{exercise}\label{ex:addsimproc}
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
  2044
  Write a simproc that replaces terms of the form @{term "t\<^sub>1 + t\<^sub>2"} by their 
130
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2045
  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
  2046
  @{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
  2047
  \end{exercise}
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2048
a21d7b300616 polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
  2049
  (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
  2050
  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
  2051
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2052
137
a9685909944d new pfd file
Christian Urban <urbanc@in.tum.de>
parents: 135
diff changeset
  2053
section {* Conversions\label{sec:conversion} *}
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
  2054
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2055
text {*
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2056
  Conversions are a thin layer on top of Isabelle's inference kernel, and can
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2057
  be viewed as a controllable, bare-bone version of Isabelle's simplifier.
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2058
  The purpose of conversions is to manipulate @{ML_type cterm}s. However,
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2059
  we will also show in this section how conversions can be applied to theorems
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2060
  and to goal states. The type of conversions is
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2061
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2062
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
  2063
ML %grayML{*type conv = cterm -> thm*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2064
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2065
text {*
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2066
  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
  2067
  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
  2068
  instance of the (meta)reflexivity theorem. For example:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2069
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2070
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2071
  "Conv.all_conv @{cterm \"Foo \<or> Bar\"}"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2072
  "Foo \<or> Bar \<equiv> Foo \<or> Bar"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2073
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2074
  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
  2075
  exception @{ML CTERM}.
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2076
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2077
  @{ML_response_fake [display,gray]
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2078
  "Conv.no_conv @{cterm True}" 
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2079
  "*** Exception- CTERM (\"no conversion\", []) raised"}
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2080
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2081
  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
  2082
  produces a meta-equation between a term and its beta-normal form. For example
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2083
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2084
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2085
  "let
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2086
  val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2087
  val two = @{cterm \"2::nat\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2088
  val ten = @{cterm \"10::nat\"}
513
f223f8223d4a updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 512
diff changeset
  2089
  val ctrm = Thm.apply (Thm.apply add two) ten
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2090
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2091
  Thm.beta_conversion true ctrm
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2092
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2093
  "((\<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
  2094
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2095
  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
  2096
  seemingly nonsensical @{term
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2097
  "2 + 10 \<equiv> 2 + (10::nat)"}. The reason is that the pretty-printer for
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2098
  @{ML_type cterm}s eta-normalises (sic) terms and therefore produces this output.
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2099
  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
  2100
  we obtain the expected result.
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2101
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2102
  @{ML_response [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2103
"let
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2104
  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
  2105
  val two = @{cterm \"2::nat\"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2106
  val ten = @{cterm \"10::nat\"}
513
f223f8223d4a updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 512
diff changeset
  2107
  val ctrm = Thm.apply (Thm.apply add two) ten
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2108
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2109
  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
  2110
end"
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 552
diff changeset
  2111
"Const (\"Pure.eq\",\<dots>) $ 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2112
  (Abs (\"x\",\<dots>,Abs (\"y\",\<dots>,\<dots>)) $\<dots>$\<dots>) $ 
418
1d1e4cda8c54 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 416
diff changeset
  2113
    (Const (\"Groups.plus_class.plus\",\<dots>) $ \<dots> $ \<dots>)"} 
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2114
512
231ec0c45bff remerged
Christian Urban <urbanc@in.tum.de>
parents: 511
diff changeset
  2115
or in the pretty-printed form
509
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2116
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2117
  @{ML_response_fake [display,gray]
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2118
"let
512
231ec0c45bff remerged
Christian Urban <urbanc@in.tum.de>
parents: 511
diff changeset
  2119
   val add = @{cterm \"\<lambda>x y. x + (y::nat)\"}
231ec0c45bff remerged
Christian Urban <urbanc@in.tum.de>
parents: 511
diff changeset
  2120
   val two = @{cterm \"2::nat\"}
231ec0c45bff remerged
Christian Urban <urbanc@in.tum.de>
parents: 511
diff changeset
  2121
   val ten = @{cterm \"10::nat\"}
513
f223f8223d4a updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 512
diff changeset
  2122
   val ctrm = Thm.apply (Thm.apply add two) ten
512
231ec0c45bff remerged
Christian Urban <urbanc@in.tum.de>
parents: 511
diff changeset
  2123
   val ctxt = @{context}
231ec0c45bff remerged
Christian Urban <urbanc@in.tum.de>
parents: 511
diff changeset
  2124
     |> Config.put eta_contract false 
231ec0c45bff remerged
Christian Urban <urbanc@in.tum.de>
parents: 511
diff changeset
  2125
     |> Config.put show_abbrevs false 
509
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2126
in
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2127
  Thm.prop_of (Thm.beta_conversion true ctrm)
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2128
  |> pretty_term ctxt
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2129
  |> pwriteln
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2130
end"
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2131
"(\<lambda>x y. x + y) 2 10 \<equiv> 2 + 10"}
dcefee89bf62 updated part about beta in Tactiacal section
Christian Urban <urbanc@in.tum.de>
parents: 506
diff changeset
  2132
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2133
  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
  2134
  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
  2135
  @{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
  2136
  on the outer-most level. 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2137
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2138
  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
  2139
  @{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
  2140
  @{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
  2141
  argument. Suppose the following meta-equation.
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2142
  
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2143
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2144
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  2145
lemma true_conj1: 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  2146
  shows "True \<and> P \<equiv> P" by simp
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2147
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2148
text {* 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2149
  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
  2150
  to @{term "Foo \<longrightarrow> Bar"}. The code is as follows.
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2151
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2152
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2153
"let 
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  2154
  val ctrm = @{cterm \"True \<and> (Foo \<longrightarrow> Bar)\"}
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2155
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2156
  Conv.rewr_conv @{thm true_conj1} ctrm
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2157
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2158
  "True \<and> (Foo \<longrightarrow> Bar) \<equiv> Foo \<longrightarrow> Bar"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2159
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2160
  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
  2161
  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
  2162
  exactly the 
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2163
  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
  2164
  the exception @{ML CTERM}.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2165
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2166
  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
  2167
  combining several conversions into one. For this you can use conversion
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  2168
  combinators. The simplest conversion combinator is @{ML_ind then_conv in Conv}, 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2169
  which applies one conversion after another. For example
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2170
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2171
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2172
"let
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2173
  val conv1 = Thm.beta_conversion false
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2174
  val conv2 = Conv.rewr_conv @{thm true_conj1}
513
f223f8223d4a updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 512
diff changeset
  2175
  val ctrm = Thm.apply @{cterm \"\<lambda>x. x \<and> False\"} @{cterm \"True\"}
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2176
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2177
  (conv1 then_conv conv2) ctrm
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2178
end"
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2179
  "(\<lambda>x. x \<and> False) True \<equiv> False"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2180
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2181
  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
  2182
  @{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
  2183
  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
  2184
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  2185
  The conversion combinator @{ML_ind else_conv in Conv} tries out the 
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2186
  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
  2187
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2188
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2189
"let
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2190
  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
  2191
  val ctrm1 = @{cterm \"True \<and> Q\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2192
  val ctrm2 = @{cterm \"P \<or> (True \<and> Q)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2193
in
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2194
  (conv ctrm1, conv ctrm2)
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2195
end"
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2196
"(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
  2197
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2198
  Here the conversion @{thm [source] true_conj1} only applies
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2199
  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
  2200
  does not fail, however, because the combinator @{ML else_conv in Conv} will then 
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2201
  try out @{ML all_conv in Conv}, which always succeeds. The same
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2202
  behaviour can also be achieved with conversion combinator @{ML_ind try_conv in Conv}.
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2203
  For example
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2204
  
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2205
  @{ML_response_fake [display,gray]
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2206
"let
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2207
  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
  2208
  val ctrm = @{cterm \"True \<or> P\"}
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2209
in
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2210
  conv ctrm
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2211
end"
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2212
  "True \<or> P \<equiv> True \<or> P"}
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2213
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2214
  Rewriting with more than one theorem can be done using the function 
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2215
  @{ML_ind rewrs_conv in Conv} from the structure @{ML_struct Conv}.
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2216
  
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2217
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  2218
  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
  2219
  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
  2220
  only apply to the outer-most level of a @{ML_type cterm}. In what follows we
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  2221
  will lift this restriction. The combinators @{ML_ind fun_conv in Conv} 
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  2222
  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
  2223
  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
  2224
  For example
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2225
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2226
  @{ML_response_fake [display,gray]
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2227
"let 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2228
  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
  2229
  val ctrm = @{cterm \"P \<or> (True \<and> Q)\"}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2230
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2231
  conv ctrm
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2232
end"
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2233
"P \<or> (True \<and> Q) \<equiv> P \<or> Q"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2234
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2235
  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
  2236
  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
  2237
  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
  2238
  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
  2239
  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
  2240
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
  2241
  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
  2242
  abstraction. For example:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2243
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2244
  @{ML_response_fake [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2245
"let 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2246
  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
  2247
  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
  2248
in
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2249
  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
  2250
end"
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2251
  "\<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
  2252
  
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2253
  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
  2254
  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
  2255
  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
  2256
  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
  2257
  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
  2258
  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
  2259
  ``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
  2260
  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
  2261
  to both branches.
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2262
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
  2263
  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
  2264
  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
  2265
  in all possible positions.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2266
*}
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2267
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2268
ML %linenosgray{*fun true_conj1_conv ctxt ctrm =
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2269
  case (Thm.term_of ctrm) of
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2270
    @{term "(\<and>)"} $ @{term True} $ _ => 
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2271
      (Conv.arg_conv (true_conj1_conv ctxt) then_conv
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2272
         Conv.rewr_conv @{thm true_conj1}) ctrm
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2273
  | _ $ _ => Conv.comb_conv (true_conj1_conv ctxt) ctrm
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2274
  | Abs _ => Conv.abs_conv (fn (_, ctxt) => true_conj1_conv ctxt) ctxt ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2275
  | _ => Conv.all_conv ctrm*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2276
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2277
text {*
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2278
  This function ``fires'' if the term is of the form @{text "(True \<and> \<dots>)"}. 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2279
  It descends under applications (Line 6) and abstractions 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2280
  (Line 7); otherwise it leaves the term unchanged (Line 8). In Line 2
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 148
diff changeset
  2281
  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
  2282
  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
  2283
  conversion in action, consider the following example:
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2284
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2285
@{ML_response_fake [display,gray]
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2286
"let
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2287
  val conv = true_conj1_conv @{context}
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2288
  val ctrm = @{cterm \"distinct [1::nat, x] \<longrightarrow> True \<and> 1 \<noteq> x\"}
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2289
in
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2290
  conv ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2291
end"
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2292
  "distinct [1, x] \<longrightarrow> True \<and> 1 \<noteq> x \<equiv> distinct [1, x] \<longrightarrow> 1 \<noteq> x"}
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2293
*}
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2294
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2295
text {*
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2296
  Conversions that traverse terms, like @{ML true_conj1_conv} above, can be
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2297
  implemented more succinctly with the combinators @{ML_ind bottom_conv in
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2298
  Conv} and @{ML_ind top_conv in Conv}. For example:
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2299
*}
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2300
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
  2301
ML %grayML{*fun true_conj1_conv ctxt =
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2302
let
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2303
  val conv = Conv.try_conv (Conv.rewr_conv @{thm true_conj1})
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2304
in
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2305
  Conv.bottom_conv (K conv) ctxt
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2306
end*}
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2307
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2308
text {*
408
Christian Urban <urbanc@in.tum.de>
parents: 406
diff changeset
  2309
  If we regard terms as trees with variables and constants on the top, then 
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2310
  @{ML bottom_conv in Conv} traverses first the the term and
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2311
  on the ``way down'' applies the conversion, whereas @{ML top_conv in
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2312
  Conv} applies the conversion on the ``way up''. To see the difference, 
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2313
  assume the following two meta-equations
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2314
*}
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2315
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2316
lemma conj_assoc:
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2317
  fixes A B C::bool
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2318
  shows "A \<and> (B \<and> C) \<equiv> (A \<and> B) \<and> C"
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2319
  and   "(A \<and> B) \<and> C \<equiv> A \<and> (B \<and> C)"
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2320
by simp_all
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2321
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2322
text {*
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2323
  and the @{ML_type cterm} @{text "(a \<and> (b \<and> c)) \<and> d"}. Here you should
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2324
  pause for a moment to be convinced that rewriting top-down and bottom-up 
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2325
  according to the two meta-equations produces two results. Below these
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2326
  two results are calculated. 
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2327
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2328
  @{ML_response_fake [display, gray]
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2329
  "let
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2330
  val ctxt = @{context}
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2331
  val conv = Conv.try_conv (Conv.rewrs_conv @{thms conj_assoc})
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2332
  val conv_top = Conv.top_conv (K conv) ctxt
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2333
  val conv_bot = Conv.bottom_conv (K conv) ctxt
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2334
  val ctrm = @{cterm \"(a \<and> (b \<and> c)) \<and> d\"}
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2335
in
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2336
  (conv_top ctrm, conv_bot ctrm)
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2337
end"
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2338
  "(\"(a \<and> (b \<and> c)) \<and> d \<equiv> a \<and> (b \<and> (c \<and> d))\",
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2339
 \"(a \<and> (b \<and> c)) \<and> d \<equiv> (a \<and> b) \<and> (c \<and> d)\")"}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2340
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2341
  To see how much control you have over rewriting with conversions, let us 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2342
  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
  2343
  @{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
  2344
  the conversion should be as follows.
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2345
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2346
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
  2347
ML %grayML{*fun if_true1_simple_conv ctxt ctrm =
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2348
  case Thm.term_of ctrm of
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2349
    Const (@{const_name If}, _) $ _ =>
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2350
      Conv.arg_conv (true_conj1_conv ctxt) ctrm
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2351
  | _ => Conv.no_conv ctrm 
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2352
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2353
val if_true1_conv = Conv.top_sweep_conv if_true1_simple_conv*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2354
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2355
text {*
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2356
  In the first function we only treat the top-most redex and also only the
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2357
  success case. As default we return @{ML no_conv in Conv}.  To apply this
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2358
  ``simple'' conversion inside a term, we use in the last line the combinator @{ML_ind
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2359
  top_sweep_conv in Conv}, which traverses the term starting from the
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2360
  root and applies it to all the redexes on the way, but stops in each branch as
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2361
  soon as it found one redex. Here is an example for this conversion:
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2362
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2363
145
f1ba430a5e7d some very slight polishing on the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 142
diff changeset
  2364
  @{ML_response_fake [display,gray]
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2365
"let
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2366
  val ctrm = @{cterm \"if P (True \<and> 1 \<noteq> (2::nat)) 
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2367
                        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
  2368
in
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2369
  if_true1_conv @{context} ctrm
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2370
end"
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2371
"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
  2372
\<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
  2373
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2374
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2375
text {*
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2376
  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
  2377
  also work on theorems using the function @{ML_ind  fconv_rule in Conv}. As an example, 
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2378
  consider again the conversion @{ML true_conj1_conv} and the lemma:
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2379
*}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2380
362
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  2381
lemma foo_test: 
a5e7ab090abf added something about manual instantiations of theorems
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
  2382
  shows "P \<or> (True \<and> \<not>P)" by simp
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2383
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2384
text {*
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2385
  Using the conversion you can transform this theorem into a 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2386
  new theorem as follows
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2387
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2388
  @{ML_response_fake [display,gray]
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2389
"let
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2390
  val conv = Conv.fconv_rule (true_conj1_conv @{context}) 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2391
  val thm = @{thm foo_test}
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2392
in
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2393
  conv thm
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2394
end" 
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2395
  "?P \<or> \<not> ?P"}
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2396
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2397
  Finally, Isabelle provides function @{ML_ind CONVERSION in Tactical} 
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2398
  for turning conversions into tactics. This needs some predefined conversion 
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2399
  combinators that traverse a goal
410
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
  2400
  state and can selectively apply the conversion. The combinators for 
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
  2401
  the goal state are: 
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2402
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2403
  \begin{itemize}
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  2404
  \item @{ML_ind params_conv in Conv} for converting under parameters
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2405
  (i.e.~where a goal state is of the form @{text "\<And>x. P x \<Longrightarrow> Q x"})
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2406
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2407
  \item @{ML_ind prems_conv in Conv} for applying a conversion to 
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2408
  premises of a goal state, and
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2409
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  2410
  \item @{ML_ind concl_conv in Conv} for applying a conversion to the
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2411
  conclusion of a goal state.
291
077c764c8d8b polished the section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 289
diff changeset
  2412
  \end{itemize}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2413
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2414
  Assume we want to apply @{ML true_conj1_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
  2415
  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
  2416
  Here is a tactic doing exactly that:
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2417
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2418
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 515
diff changeset
  2419
ML %grayML{*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
  2420
  CONVERSION 
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 184
diff changeset
  2421
    (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
  2422
       (Conv.prems_conv ~1 (if_true1_conv ctxt) then_conv
405
f8d020bbc2c0 improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents: 404
diff changeset
  2423
          Conv.concl_conv ~1 (true_conj1_conv ctxt))) ctxt)*}
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2424
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2425
text {* 
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2426
  We call the conversions with the argument @{ML "~1"}. By this we 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2427
  analyse all parameters, all premises and the conclusion of a goal
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2428
  state. If we call @{ML concl_conv in Conv} with 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2429
  a non-negative number, say @{text n}, then this conversions will 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2430
  skip the first @{text n} premises and applies the conversion to the 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2431
  ``rest'' (similar for parameters and conclusions). To test the 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2432
  tactic, consider the proof
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2433
*}
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2434
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2435
lemma
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2436
  "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
  2437
  (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
  2438
apply(tactic {* true1_tac @{context} 1 *})
147
6dafb0815ae6 more polishing of the conversion section
Christian Urban <urbanc@in.tum.de>
parents: 146
diff changeset
  2439
txt {* where the tactic yields the goal state
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2440
  \begin{isabelle}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 174
diff changeset
  2441
  @{subgoals [display]}
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2442
  \end{isabelle}*}
142
c06885c36575 Polished conversion section.
boehmes
parents: 141
diff changeset
  2443
(*<*)oops(*>*)
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2444
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2445
text {*
148
84d1392186d3 deleted Ad-hoc recipe
Christian Urban <urbanc@in.tum.de>
parents: 147
diff changeset
  2446
  As you can see, the premises are rewritten according to @{ML if_true1_conv}, while
410
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
  2447
  the conclusion according to @{ML true_conj1_conv}. If we only have one
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
  2448
  conversion that should be uniformly applied to the whole goal state, we
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2449
  can simplify @{ML true1_tac} using @{ML_ind top_conv in Conv}.
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2450
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2451
  Conversions are also be helpful for constructing meta-equality theorems.
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2452
  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
  2453
  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
  2454
*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2455
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2456
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
  2457
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
  2458
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2459
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
  2460
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
  2461
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2462
text {*
335
163ac0662211 reorganised the certified terms section; tuned
Christian Urban <urbanc@in.tum.de>
parents: 334
diff changeset
  2463
  Although both definitions define the same function, the theorems @{thm
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2464
  [source] id1_def} and @{thm [source] id2_def} are different meta-equations. However it is
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2465
  easy to transform one into the other. The function @{ML_ind abs_def
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2466
  in Drule} from the structure @{ML_struct Drule} can automatically transform 
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2467
  theorem @{thm [source] id1_def}
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2468
  into @{thm [source] id2_def} by abstracting all variables on the 
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2469
  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
  2470
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2471
  @{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
  2472
  "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
  2473
  "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
  2474
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2475
  Unfortunately, Isabelle has no built-in function that transforms the
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2476
  theorems in the other direction. We can implement one using
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2477
  conversions. The feature of conversions we are using is that if we apply a
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2478
  @{ML_type cterm} to a conversion we obtain a meta-equality theorem (recall
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2479
  that the type of conversions is an abbreviation for 
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2480
  @{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
  2481
*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2482
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2483
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
  2484
let
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2485
  val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of def)
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2486
  val xs = strip_abs_vars (Thm.term_of rhs)
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2487
  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
  2488
  
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  2489
  val cxs = map (Thm.cterm_of ctxt' o Free) xs
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2490
  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
  2491
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2492
  fun get_conv [] = Conv.rewr_conv def
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2493
    | 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
  2494
in
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2495
  get_conv xs new_lhs |>  
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
  2496
  singleton (Proof_Context.export ctxt' ctxt)
332
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2497
end*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2498
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2499
text {*
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2500
  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
  2501
  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
  2502
  assumption in @{ML unabs_def} is that the right-hand side is an
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2503
  abstraction. We compute the possibly empty list of abstracted variables in
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 368
diff changeset
  2504
  Line 4 using the function @{ML_ind strip_abs_vars in Term}. For this we have to
412
73f716b9201a polised
Christian Urban <urbanc@in.tum.de>
parents: 411
diff changeset
  2505
  first transform the @{ML_type cterm} into a @{ML_type term}. In Line 5 we
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2506
  import these variables into the context @{text "ctxt'"}, in order to
334
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2507
  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
  2508
  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
  2509
  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
  2510
  conversion according to the length of the list of (abstracted) variables. If
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2511
  there are none, then we just return the conversion corresponding to the
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2512
  original definition. If there are variables, then we have to prefix this
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2513
  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
  2514
  apply the new left-hand side to the generated conversion and obtain the the
Christian Urban <urbanc@in.tum.de>
parents: 333
diff changeset
  2515
  theorem we want to construct. As mentioned above, in Line 15 we only have to
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2516
  export the context @{text "ctxt'"} back to @{text "ctxt"} in order to 
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2517
  produce meta-variables for 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
  2518
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2519
  @{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
  2520
  "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
  2521
  "id2 ?x1 \<equiv> ?x1"}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2522
*}
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2523
6fba3a3e80a3 added the function unabs_def in the conversions section
Christian Urban <urbanc@in.tum.de>
parents: 331
diff changeset
  2524
text {*
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2525
  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
  2526
  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
  2527
  often much less efficient than the simplifier. The advantage of conversions, 
406
f399b6855546 implemented suggestion from Sascha
Christian Urban <urbanc@in.tum.de>
parents: 405
diff changeset
  2528
  however, is 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
  2529
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
  2530
  \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
  2531
  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
  2532
  \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
  2533
  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
  2534
  \end{exercise}
8084c353d196 added material to the endless story of the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 151
diff changeset
  2535
172
ec47352e99c2 improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents: 170
diff changeset
  2536
  \begin{exercise}\label{ex:compare}
174
a29b81d4fa88 some minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
  2537
  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
  2538
  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
  2539
  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
  2540
  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
  2541
  \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
  2542
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2543
  \begin{readmore}
424
5e0a2b50707e updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 422
diff changeset
  2544
  See @{ML_file "Pure/conv.ML"}
384
Christian Urban <urbanc@in.tum.de>
parents: 378
diff changeset
  2545
  for more information about conversion combinators. 
243
6e0f56764ff8 slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 241
diff changeset
  2546
  Some basic conversions are defined in  @{ML_file "Pure/thm.ML"},
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 456
diff changeset
  2547
  @{ML_file "Pure/drule.ML"} and @{ML_file "Pure/raw_simplifier.ML"}.
146
4aa8a80e37ff some polishing about conversions
Christian Urban <urbanc@in.tum.de>
parents: 145
diff changeset
  2548
  \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
  2549
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2550
*}
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2551
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
  2552
text {*
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
  2553
  (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
  2554
  are of any use/efficient)
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
  2555
*}
135
8c31b729a5df First draft for conversions subsection.
boehmes
parents: 134
diff changeset
  2556
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
  2557
216
fcedd5bd6a35 added a declaration section (for Amine)
Christian Urban <urbanc@in.tum.de>
parents: 214
diff changeset
  2558
section {* Declarations (TBD) *}
fcedd5bd6a35 added a declaration section (for Amine)
Christian Urban <urbanc@in.tum.de>
parents: 214
diff changeset
  2559
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 489
diff changeset
  2560
section {* Structured Proofs\label{sec:structured} (TBD) *}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2561
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2562
text {* TBD *}
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
  2563
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2564
lemma True
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2565
proof
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2566
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2567
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2568
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2569
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2570
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2571
    then have "A & B" ..
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2572
    then have C by (rule r)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2573
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2574
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2575
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2576
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2577
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2578
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2579
    note conjI [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2580
    note r [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2581
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2582
oops
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2583
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2584
ML {* 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2585
  val ctxt0 = @{context};
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2586
  val ctxt = ctxt0;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2587
  val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt;
217
75154f4d4e2f used antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 216
diff changeset
  2588
  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
  2589
  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
  2590
  val this = [@{thm conjI} OF this]; 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2591
  val this = r OF this;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2592
  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
  2593
  val this = Variable.export ctxt ctxt0 [this] 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
  2594
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2595
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2596
section {* Summary *}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2597
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  2598
text {*
368
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2599
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2600
  \begin{conventions}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2601
  \begin{itemize}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2602
  \item Reference theorems with the antiquotation @{text "@{thm \<dots>}"}.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2603
  \item If a tactic is supposed to fail, return the empty sequence.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2604
  \item If you apply a tactic to a sequence of subgoals, apply it 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2605
  in reverse order (i.e.~start with the last subgoal). 
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2606
  \item Use simpsets that are as small as possible.
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2607
  \end{itemize}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2608
  \end{conventions}
b1a458a03a8e new parts in the tactical section
Christian Urban <urbanc@in.tum.de>
parents: 363
diff changeset
  2609
363
Christian Urban <urbanc@in.tum.de>
parents: 362
diff changeset
  2610
*}
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
  2611
139
ed1eb9cb2533 Improved conversion subsection.
boehmes
parents: 135
diff changeset
  2612
end