CookBook/Tactical.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 05 Feb 2009 22:44:11 +0000
changeset 100 dd8eebae11ec
parent 99 de625e5f6a36
child 102 5e309df58557
permissions -rw-r--r--
polished
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Tactical
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
     2
imports Base FirstSteps
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
chapter {* Tactical Reasoning\label{chp:tactical} *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  The main reason for descending to the ML-level of Isabelle is to be able to
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  implement automatic proof procedures. Such proof procedures usually lessen
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  considerably the burden of manual reasoning, for example, when introducing
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  new definitions. These proof procedures are centred around refining a goal
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  state using tactics. This is similar to the @{text apply}-style reasoning at
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  the user level, where goals are modified in a sequence of proof steps until
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    15
  all of them are solved. However, there are also more structured operations
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    16
  that help with handling of variables and assumptions.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
section {* Tactical Reasoning *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  To see how tactics work, let us first transcribe a simple @{text apply}-style proof 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  into ML. Consider the following proof.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
lemma disj_swap: "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
apply(erule disjE)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
apply(rule disjI2)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
apply(rule disjI1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
apply(assumption)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  This proof translates to the following ML-code.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
@{ML_response_fake [display,gray]
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
"let
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  val goal = @{prop \"P \<or> Q \<Longrightarrow> Q \<or> P\"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
in
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  Goal.prove ctxt [\"P\", \"Q\"] [] goal 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
   (fn _ => 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
      etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
      THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
      THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      THEN atac 1)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
end" "?P \<or> ?Q \<Longrightarrow> ?Q \<or> ?P"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  To start the proof, the function @{ML "Goal.prove"}~@{text "ctxt xs As C
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    52
  tac"} sets up a goal state for proving the goal @{text C} 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    53
  (that is @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"} in the proof at hand) under the
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    54
  assumptions @{text As} (happens to be empty) with the variables
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  @{text xs} that will be generalised once the goal is proved (in our case
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  @{text P} and @{text Q}). The @{text "tac"} is the tactic that proves the goal;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  it can make use of the local assumptions (there are none in this example). 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  The functions @{ML etac}, @{ML rtac} and @{ML atac} correspond to 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  @{text erule}, @{text rule} and @{text assumption}, respectively. 
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    60
  The operator @{ML THEN} strings the tactics together.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  \begin{readmore}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    63
  To learn more about the function @{ML Goal.prove} see \isccite{sec:results}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    64
  and the file @{ML_file "Pure/goal.ML"}. For more information about the
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    65
  internals of goals see \isccite{sec:tactical-goals}.  See @{ML_file
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    66
  "Pure/tactic.ML"} and @{ML_file "Pure/tctical.ML"} for the code of basic
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    67
  tactics and tactic combinators; see also Chapters 3 and 4 in the old
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    68
  Isabelle Reference Manual.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  Note that we used antiquotations for referencing the theorems. We could also
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  just have written @{ML "etac disjE 1"} and so on, but this is considered bad
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  style. The reason is that the binding for @{ML disjE} can be re-assigned by 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  the user and thus one does not have complete control over which theorem is
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  actually applied. This problem is nicely prevented by using antiquotations, 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  because then the theorems are fixed statically at compile-time.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  During the development of automatic proof procedures, it will often be necessary 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  to test a tactic on examples. This can be conveniently 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  done with the command \isacommand{apply}@{text "(tactic \<verbopen> \<dots> \<verbclose>)"}. 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
  Consider the following sequence of tactics
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
ML{*val foo_tac = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
      (etac @{thm disjE} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
       THEN rtac @{thm disjI2} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
       THEN atac 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
       THEN rtac @{thm disjI1} 1
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
       THEN atac 1)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
text {* and the Isabelle proof: *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
lemma "P \<or> Q \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
apply(tactic {* foo_tac *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    98
  By using @{text "tactic \<verbopen> \<dots> \<verbclose>"} in the apply-step,
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
    99
  you can call @{ML foo_tac} or any function that returns a tactic from the
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   100
  user level of Isabelle.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
  As can be seen, each tactic in @{ML foo_tac} has a hard-coded number that
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  stands for the subgoal analysed by the tactic. In our case, we only focus on the first
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  subgoal. This is sometimes wanted, but usually not. To avoid the explicit numbering in the 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  tactic, you can write
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
ML{*val foo_tac' = 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
      (etac @{thm disjE} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
       THEN' rtac @{thm disjI2} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
       THEN' atac 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
       THEN' rtac @{thm disjI1} 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
       THEN' atac)*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
text {* 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  and then give the number for the subgoal explicitly when the tactic is
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   117
  called. For every operator that combines tactics, such a primed version 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   118
  exists. So in the next proof we can first discharge the second subgoal,
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  and after that the first.
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
   and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
apply(tactic {* foo_tac' 2 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   129
  (FIXME: maybe add something about how each goal state is interpreted
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   130
  as a theorem)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   131
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   132
  The tactic @{ML foo_tac} (and @{ML foo_tac'}) are very specific for
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   133
  analysing goals only of the form @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}. If the goal is not
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   134
  of this form, then @{ML foo_tac} throws the error message:
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   135
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   136
  \begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   137
  @{text "*** empty result sequence -- proof command failed"}\\
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   138
  @{text "*** At command \"apply\"."}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   139
  \end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   140
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   141
  Meaning the tactic failed. The reason for this error message is that tactics 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   142
  are functions that map a goal state to a (lazy) sequence of successor states, 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   143
  hence the type of a tactic is
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  @{text [display, gray] "type tactic = thm -> thm Seq.seq"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   147
  It is custom that if a tactic fails, it should return the empty sequence: 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   148
  tactics should not raise exceptions willy-nilly.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   150
  The lazy list of possible successor states shows through to the user-level 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   151
  of Isabelle when using the command \isacommand{back}. For instance in the 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   152
  following proof, there are two possibilities for how to apply 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   153
  @{ML foo_tac'}.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
apply(tactic {* foo_tac' 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
back
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   162
  By using \isacommand{back}, we construct the proof that uses the
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   163
  second assumption to complete the proof. In more interesting 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   164
  situations, different possibilities can lead to different proofs.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   165
  
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  See @{ML_file "Pure/General/seq.ML"} for the implementation of lazy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  sequences. However in day-to-day Isabelle programming, one rarely 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  constructs sequences explicitly, but uses the predefined functions
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   170
  instead.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  \end{readmore}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
section {* Basic Tactics *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   177
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   178
  As seen above, the function @{ML atac} corresponds to the assumption tactic.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   179
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   180
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   181
lemma shows "P \<Longrightarrow> P"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
apply(tactic {* atac 1 *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
done
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   185
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   186
  Similarly, @{ML rtac}, @{ML dtac}, @{ML etac} and @{ML ftac} correspond
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   187
  to @{ML_text rule}, @{ML_text drule}, @{ML_text erule} and @{ML_text frule}, 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   188
  respectively. Below are three examples with the resulting goal state. How
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   189
  they work should therefore be self-explanatory.  
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   190
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   191
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   192
lemma shows "P \<and> Q"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
apply(tactic {* rtac @{thm conjI} 1 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   194
txt{*@{subgoals [display]}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   197
lemma shows "P \<and> Q \<Longrightarrow> False"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
apply(tactic {* etac @{thm conjE} 1 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   199
txt{*@{subgoals [display]}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
lemma shows "False \<and> True \<Longrightarrow> False"
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
apply(tactic {* dtac @{thm conjunct2} 1 *})
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   204
txt{*@{subgoals [display]}*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   208
  As mentioned above, most basic tactics take a number as argument, which
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   209
  addresses to subgoal they are analysing.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   210
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   211
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   212
lemma shows "Foo" and "P \<and> Q"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   213
apply(tactic {* rtac @{thm conjI} 2 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   214
txt {*@{subgoals [display]}*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   215
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   216
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   217
text {* 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   218
  Corresponding to @{ML rtac}, there is also the tactic @{ML resolve_tac}, which
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   219
  however expects a list of theorems as arguments. From this list it will apply with 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   220
  the first applicable theorem (later theorems that are also applicable can be
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   221
  explored via the lazy sequences mechanism). Given the abbreviation
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   224
ML{*val resolve_tac_xmp = resolve_tac [@{thm impI}, @{thm conjI}]*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   225
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   226
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   227
  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
   228
  implication is analysed and then an outermost conjunction.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   229
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   230
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   231
lemma shows "C \<longrightarrow> (A \<and> B)" and "(A \<longrightarrow> B) \<and> C"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   232
apply(tactic {* resolve_tac_xmp 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   233
apply(tactic {* resolve_tac_xmp 2 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   234
txt{* @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   235
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   236
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   237
text {* 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   238
  Similarly versions exists for @{ML atac} (@{ML assume_tac}), @{ML etac} 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   239
  (@{ML eresolve_tac}) and so on.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   240
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   241
  The tactic @{ML print_tac} is useful for low-level debugging of tactics: it
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   242
  prints out a message and the current goal state.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   243
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   244
 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   245
lemma shows "False \<Longrightarrow> True"
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
apply(tactic {* print_tac "foo message" *})
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   249
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   250
  Also for useful for debugging, but not exclusively, are the tactics @{ML all_tac} and
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   251
  @{ML no_tac}. The former always succeeds (but does not change the goal state), and
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   252
  the latter always fails.
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   253
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   254
  (FIXME explain RS MRS)
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   255
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   256
  Often proofs involve elaborate operations on assumptions and variables
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   257
  @{text "\<And>"}-quantified in the goal state. To do such operations on the ML-level 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   258
  using the basic tactics, is very unwieldy and brittle. Some convenience and
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   259
  safety is provided by the tactic @{ML SUBPROOF}. This tactic fixes the parameters 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   260
  and binds the various components of a proof state into a record. 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   261
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   262
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   263
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   264
\begin{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   265
\begin{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   266
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   267
ML{*fun sp_tac {prems, params, asms, concl, context, schematics} = 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   268
  let 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   269
    val str_of_params = str_of_cterms context params
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   270
    val str_of_asms = str_of_cterms context asms
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   271
    val str_of_concl = str_of_cterm context concl
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   272
    val str_of_prems = str_of_thms context prems   
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   273
    val str_of_schms = str_of_cterms context (snd schematics)    
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   274
 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   275
    val _ = (warning ("params: " ^ str_of_params);
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   276
             warning ("schematics: " ^ str_of_schms);
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   277
             warning ("asms: " ^ str_of_asms);
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   278
             warning ("concl: " ^ str_of_concl);
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   279
             warning ("prems: " ^ str_of_prems))
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   280
  in
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   281
    no_tac 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   282
  end*}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   283
text_raw{*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   284
\end{isabelle}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   285
\caption{A function that prints out the various parameters provided by the tactic
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   286
 @{ML SUBPROOF}. It uses the functions extracting strings from @{ML_type cterm}s 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   287
  and @{ML_type thm}s, which are defined in Section~\ref{sec:printing}.\label{fig:sptac}}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   288
\end{figure}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   289
*}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   290
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   291
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   292
  To see what happens, assume the function defined in Figure~\ref{fig:sptac}, which
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   293
  takes a record as argument and just prints out the content of this record (using the 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   294
  string transformation functions defined in Section~\ref{sec:printing}). Consider
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   295
  now the proof
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   296
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   297
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   298
lemma shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x"
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   299
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   300
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   301
txt {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   302
  which yields the printout:
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   303
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   304
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   305
  \begin{tabular}{ll}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   306
  params:     & @{term x}, @{term y}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   307
  schematics: & @{term z}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   308
  asms:       & @{term "A x y"}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   309
  concl:      & @{term "B y x \<longrightarrow> C (z y) x"}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   310
  prems:      & @{term "A x y"}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   311
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   312
  \end{quote}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   313
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   314
  Note in the actual output the brown colour of the variables @{term x} and 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   315
  @{term y}. Although parameters in the original goal, they are fixed inside
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   316
  the subproof. Similarly the schematic variable @{term z}. The assumption 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   317
  is bound once as @{ML_type cterm} to the record-variable @{text asms} and 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   318
  another time 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
   319
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   320
  Notice also that we had to append @{text "?"} to \isacommand{apply}. The 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   321
  reason is that @{ML SUBPROOF} normally expects that the subgoal is solved completely.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   322
  Since in the function @{ML sp_tac} we returned the tactic @{ML no_tac}, the subproof
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   323
  obviously fails. The question-mark allows us to recover from this failure
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   324
  in a graceful manner so that the warning messages are not overwritten
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   325
  by the error message.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   326
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   327
  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
   328
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   329
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   330
apply(rule impI)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   331
apply(tactic {* SUBPROOF sp_tac @{context} 1 *})?
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   332
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   333
txt {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   334
  then @{ML SUBPROOF} prints out 
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   335
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   336
  \begin{quote}\small
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   337
  \begin{tabular}{ll}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   338
  params:     & @{term x}, @{term y}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   339
  schematics: & @{term z}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   340
  asms:       & @{term "A x y"}, @{term "B y x"}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   341
  concl:      & @{term "C (z y) x"}\\
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   342
  prems:      & @{term "A x y"}, @{term "B y x"}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   343
  \end{tabular}
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   344
  \end{quote}
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   345
*}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   346
(*<*)oops(*>*)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   347
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   348
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   349
  where we now also have @{term "B y x"} as assumption.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   350
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   351
  One convenience of @{ML SUBPROOF} is that we can apply assumption
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   352
  using the usual tactics, because the parameter @{text prems} 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   353
  contains the assumptions as theorems. With this we can easily 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   354
  implement a tactic that almost behaves like @{ML atac}:
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   355
*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   356
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   357
lemma shows "\<And>x y. \<lbrakk>B x y; A x y; C x y\<rbrakk> \<Longrightarrow> A x y"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   358
apply(tactic 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   359
       {* SUBPROOF (fn {prems, ...} => resolve_tac prems 1) @{context} 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   360
txt{* yields
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   361
      @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   362
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   363
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   364
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   366
  The restriction in this tactic is that it cannot instantiate any
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   367
  schematic variables. This might be seen as a defect, but is actually
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   368
  an advantage in the situations for which @{ML SUBPROOF} was designed: 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   369
  the reason is that instantiation of schematic variables can potentially 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   370
  has affect several goals and can render them unprovable.  
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   371
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   372
  A similar but less powerful function is @{ML SUBGOAL}. It allows you to 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   373
  inspect a subgoal specified by a number. 
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   376
ML %linenumbers{*fun select_tac (t,i) =
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   377
  case t of
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   378
     @{term "Trueprop"} $ t' => select_tac (t',i)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   379
   | @{term "op \<and>"} $ _ $ _ => rtac @{thm conjI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   380
   | @{term "op \<longrightarrow>"} $ _ $ _ => rtac @{thm impI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   381
   | @{term "Not"} $ _ => rtac @{thm notI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   382
   | Const (@{const_name "All"}, _) $ _ => rtac @{thm allI} i
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   383
   | _ => no_tac*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   384
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   385
lemma shows "A \<and> B" "A \<longrightarrow> B" "\<forall>x. C x"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   386
apply(tactic {* SUBGOAL select_tac 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   387
apply(tactic {* SUBGOAL select_tac 3 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   388
apply(tactic {* SUBGOAL select_tac 4 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   389
txt{* @{subgoals [display]} *}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   390
(*<*)oops(*>*)
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   391
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   392
text {*
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   393
  However, this example is contrived, as there are much simpler ways
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   394
  to implement a proof procedure like the one above. They will be explained
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   395
  in the next section.
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   396
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   397
  A variant of @{ML SUBGOAL} is @{ML CSUBGOAL} which allows access to the goal
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   398
  as @{ML_type cterm} instead of a @{ML_type term}.
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   401
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
section {* Operations on Tactics *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   404
text {* @{ML THEN} *}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   406
lemma shows "(Foo \<and> Bar) \<and> False"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   407
apply(tactic {* (rtac @{thm conjI} 1) 
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   408
                 THEN (rtac @{thm conjI} 1) *})
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
txt {* @{subgoals [display]} *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   412
ML{*val orelse_xmp = (rtac @{thm disjI1} ORELSE' rtac @{thm conjI})*}
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   413
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   414
lemma shows "True \<and> False" and "Foo \<or> Bar"
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   415
apply(tactic {* orelse_xmp 1 *})
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   416
apply(tactic {* orelse_xmp 3 *})
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
txt {* @{subgoals [display]} *}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
(*<*)oops(*>*)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
text {*
99
de625e5f6a36 polished
Christian Urban <urbanc@in.tum.de>
parents: 95
diff changeset
   422
  @{ML EVERY} @{ML REPEAT} @{ML DETERM}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
  @{ML rewrite_goals_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
  @{ML cut_facts_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
  @{ML ObjectLogic.full_atomize_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
  @{ML ObjectLogic.rulify_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
  @{ML resolve_tac}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
text {*
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
  A goal (or goal state) is a special @{ML_type thm}, which by
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
  convention is an implication of the form:
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  @{text[display] "A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> #(C)"}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
  where @{term C} is the goal to be proved and the @{term "A\<^isub>i"} are the open 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
  subgoals. 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
  Since the goal @{term C} can potentially be an implication, there is a
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
  @{text "#"} wrapped around it, which prevents that premises are 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
  misinterpreted as open subgoals. The wrapper @{text "# :: prop \<Rightarrow>
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
  prop"} is just the identity function and used as a syntactic marker. 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
  
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
  While tactics can operate on the subgoals (the @{text "A\<^isub>i"} above), they 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
  are expected to leave the conclusion @{term C} intact, with the 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
  exception of possibly instantiating schematic variables. 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
 
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
*}
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
95
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   460
section {* Structured Proofs *}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   461
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   462
lemma True
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   463
proof
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   464
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   465
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   466
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   467
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   468
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   469
    then have "A & B" ..
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   470
    then have C by (rule r)
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   471
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   472
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   473
  {
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   474
    fix A B C
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   475
    assume r: "A & B \<Longrightarrow> C"
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   476
    assume A B
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   477
    note conjI [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   478
    note r [OF this]
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   479
  }
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   480
oops
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   481
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   482
ML {* fun prop ctxt s =
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   483
  Thm.cterm_of (ProofContext.theory_of ctxt) (Syntax.read_prop ctxt s) *}
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   484
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   485
ML {* 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   486
  val ctxt0 = @{context};
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   487
  val ctxt = ctxt0;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   488
  val (_, ctxt) = Variable.add_fixes ["A", "B", "C"] ctxt;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   489
  val ([r], ctxt) = Assumption.add_assumes [prop ctxt "A & B \<Longrightarrow> C"] ctxt;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   490
  val (this, ctxt) = Assumption.add_assumes [prop ctxt "A", prop ctxt "B"] ctxt;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   491
  val this = [@{thm conjI} OF this]; 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   492
  val this = r OF this;
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   493
  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
   494
  val this = Variable.export ctxt ctxt0 [this] 
7235374f34c8 added some preliminary notes about SUBPROOF
Christian Urban <urbanc@in.tum.de>
parents: 93
diff changeset
   495
*}
93
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
end