doc-src/IsarRef/Thy/Framework.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 12 Aug 2010 15:03:34 +0200
branchisac-from-Isabelle2009-2
changeset 37913 20e3616b2d9c
parent 36367 641a521bfc19
child 40745 515eab39b6c2
permissions -rw-r--r--
prepare reactivation of isac-update-Isa09-2
wenzelm@30042
     1
theory Framework
wenzelm@30042
     2
imports Main
wenzelm@30042
     3
begin
wenzelm@30042
     4
wenzelm@30042
     5
chapter {* The Isabelle/Isar Framework \label{ch:isar-framework} *}
wenzelm@30042
     6
wenzelm@30042
     7
text {*
wenzelm@30042
     8
  Isabelle/Isar
wenzelm@30042
     9
  \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift}
wenzelm@30042
    10
  is intended as a generic framework for developing formal
wenzelm@30042
    11
  mathematical documents with full proof checking.  Definitions and
wenzelm@30061
    12
  proofs are organized as theories.  An assembly of theory sources may
wenzelm@30042
    13
  be presented as a printed document; see also
wenzelm@30042
    14
  \chref{ch:document-prep}.
wenzelm@30042
    15
wenzelm@30042
    16
  The main objective of Isar is the design of a human-readable
wenzelm@30042
    17
  structured proof language, which is called the ``primary proof
wenzelm@30042
    18
  format'' in Isar terminology.  Such a primary proof language is
wenzelm@30042
    19
  somewhere in the middle between the extremes of primitive proof
wenzelm@30042
    20
  objects and actual natural language.  In this respect, Isar is a bit
wenzelm@30042
    21
  more formalistic than Mizar
wenzelm@30042
    22
  \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar},
wenzelm@30042
    23
  using logical symbols for certain reasoning schemes where Mizar
wenzelm@30042
    24
  would prefer English words; see \cite{Wenzel-Wiedijk:2002} for
wenzelm@30042
    25
  further comparisons of these systems.
wenzelm@30042
    26
wenzelm@30042
    27
  So Isar challenges the traditional way of recording informal proofs
wenzelm@30042
    28
  in mathematical prose, as well as the common tendency to see fully
wenzelm@30042
    29
  formal proofs directly as objects of some logical calculus (e.g.\
wenzelm@30042
    30
  @{text "\<lambda>"}-terms in a version of type theory).  In fact, Isar is
wenzelm@30042
    31
  better understood as an interpreter of a simple block-structured
wenzelm@30061
    32
  language for describing the data flow of local facts and goals,
wenzelm@30042
    33
  interspersed with occasional invocations of proof methods.
wenzelm@30042
    34
  Everything is reduced to logical inferences internally, but these
wenzelm@30042
    35
  steps are somewhat marginal compared to the overall bookkeeping of
wenzelm@30042
    36
  the interpretation process.  Thanks to careful design of the syntax
wenzelm@30042
    37
  and semantics of Isar language elements, a formal record of Isar
wenzelm@30042
    38
  instructions may later appear as an intelligible text to the
wenzelm@30042
    39
  attentive reader.
wenzelm@30042
    40
wenzelm@30042
    41
  The Isar proof language has emerged from careful analysis of some
wenzelm@30042
    42
  inherent virtues of the existing logical framework of Isabelle/Pure
wenzelm@30042
    43
  \cite{paulson-found,paulson700}, notably composition of higher-order
wenzelm@30042
    44
  natural deduction rules, which is a generalization of Gentzen's
wenzelm@30042
    45
  original calculus \cite{Gentzen:1935}.  The approach of generic
wenzelm@30042
    46
  inference systems in Pure is continued by Isar towards actual proof
wenzelm@30042
    47
  texts.
wenzelm@30042
    48
wenzelm@30042
    49
  Concrete applications require another intermediate layer: an
wenzelm@30042
    50
  object-logic.  Isabelle/HOL \cite{isa-tutorial} (simply-typed
wenzelm@30042
    51
  set-theory) is being used most of the time; Isabelle/ZF
wenzelm@30042
    52
  \cite{isabelle-ZF} is less extensively developed, although it would
wenzelm@30042
    53
  probably fit better for classical mathematics.
wenzelm@30047
    54
wenzelm@30061
    55
  \medskip In order to illustrate natural deduction in Isar, we shall
wenzelm@30061
    56
  refer to the background theory and library of Isabelle/HOL.  This
wenzelm@30061
    57
  includes common notions of predicate logic, naive set-theory etc.\
wenzelm@30061
    58
  using fairly standard mathematical notation.  From the perspective
wenzelm@30061
    59
  of generic natural deduction there is nothing special about the
wenzelm@30061
    60
  logical connectives of HOL (@{text "\<and>"}, @{text "\<or>"}, @{text "\<forall>"},
wenzelm@30061
    61
  @{text "\<exists>"}, etc.), only the resulting reasoning principles are
wenzelm@30061
    62
  relevant to the user.  There are similar rules available for
wenzelm@30061
    63
  set-theory operators (@{text "\<inter>"}, @{text "\<union>"}, @{text "\<Inter>"}, @{text
wenzelm@30061
    64
  "\<Union>"}, etc.), or any other theory developed in the library (lattice
wenzelm@30061
    65
  theory, topology etc.).
wenzelm@30047
    66
wenzelm@30047
    67
  Subsequently we briefly review fragments of Isar proof texts
wenzelm@30061
    68
  corresponding directly to such general deduction schemes.  The
wenzelm@30061
    69
  examples shall refer to set-theory, to minimize the danger of
wenzelm@30047
    70
  understanding connectives of predicate logic as something special.
wenzelm@30047
    71
wenzelm@30047
    72
  \medskip The following deduction performs @{text "\<inter>"}-introduction,
wenzelm@30047
    73
  working forwards from assumptions towards the conclusion.  We give
wenzelm@30047
    74
  both the Isar text, and depict the primitive rule involved, as
wenzelm@30061
    75
  determined by unification of the problem against rules that are
wenzelm@30061
    76
  declared in the library context.
wenzelm@30047
    77
*}
wenzelm@30047
    78
wenzelm@30047
    79
text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
wenzelm@30047
    80
wenzelm@30047
    81
(*<*)
wenzelm@36367
    82
example_proof
wenzelm@30047
    83
(*>*)
wenzelm@30047
    84
    assume "x \<in> A" and "x \<in> B"
wenzelm@30047
    85
    then have "x \<in> A \<inter> B" ..
wenzelm@30047
    86
(*<*)
wenzelm@30047
    87
qed
wenzelm@30047
    88
(*>*)
wenzelm@30047
    89
wenzelm@30047
    90
text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
wenzelm@30047
    91
wenzelm@30047
    92
text {*
wenzelm@30047
    93
  \infer{@{prop "x \<in> A \<inter> B"}}{@{prop "x \<in> A"} & @{prop "x \<in> B"}}
wenzelm@30047
    94
*}
wenzelm@30047
    95
wenzelm@30047
    96
text_raw {*\end{minipage}*}
wenzelm@30047
    97
wenzelm@30047
    98
text {*
wenzelm@30063
    99
  \medskip\noindent Note that @{command assume} augments the proof
wenzelm@30063
   100
  context, @{command then} indicates that the current fact shall be
wenzelm@30063
   101
  used in the next step, and @{command have} states an intermediate
wenzelm@30061
   102
  goal.  The two dots ``@{command ".."}'' refer to a complete proof of
wenzelm@30061
   103
  this claim, using the indicated facts and a canonical rule from the
wenzelm@30047
   104
  context.  We could have been more explicit here by spelling out the
wenzelm@30047
   105
  final proof step via the @{command "by"} command:
wenzelm@30047
   106
*}
wenzelm@30047
   107
wenzelm@30047
   108
(*<*)
wenzelm@36367
   109
example_proof
wenzelm@30047
   110
(*>*)
wenzelm@30047
   111
    assume "x \<in> A" and "x \<in> B"
wenzelm@30047
   112
    then have "x \<in> A \<inter> B" by (rule IntI)
wenzelm@30047
   113
(*<*)
wenzelm@30047
   114
qed
wenzelm@30047
   115
(*>*)
wenzelm@30047
   116
wenzelm@30047
   117
text {*
wenzelm@30047
   118
  \noindent The format of the @{text "\<inter>"}-introduction rule represents
wenzelm@30047
   119
  the most basic inference, which proceeds from given premises to a
wenzelm@30061
   120
  conclusion, without any nested proof context involved.
wenzelm@30047
   121
wenzelm@30061
   122
  The next example performs backwards introduction on @{term "\<Inter>\<A>"},
wenzelm@30061
   123
  the intersection of all sets within a given set.  This requires a
wenzelm@30061
   124
  nested proof of set membership within a local context, where @{term
wenzelm@30061
   125
  A} is an arbitrary-but-fixed member of the collection:
wenzelm@30047
   126
*}
wenzelm@30047
   127
wenzelm@30047
   128
text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
wenzelm@30047
   129
wenzelm@30047
   130
(*<*)
wenzelm@36367
   131
example_proof
wenzelm@30047
   132
(*>*)
wenzelm@30047
   133
    have "x \<in> \<Inter>\<A>"
wenzelm@30047
   134
    proof
wenzelm@30047
   135
      fix A
wenzelm@30047
   136
      assume "A \<in> \<A>"
wenzelm@30059
   137
      show "x \<in> A" sorry %noproof
wenzelm@30047
   138
    qed
wenzelm@30047
   139
(*<*)
wenzelm@30047
   140
qed
wenzelm@30047
   141
(*>*)
wenzelm@30047
   142
wenzelm@30047
   143
text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
wenzelm@30047
   144
wenzelm@30047
   145
text {*
wenzelm@30047
   146
  \infer{@{prop "x \<in> \<Inter>\<A>"}}{\infer*{@{prop "x \<in> A"}}{@{text "[A][A \<in> \<A>]"}}}
wenzelm@30047
   147
*}
wenzelm@30047
   148
wenzelm@30047
   149
text_raw {*\end{minipage}*}
wenzelm@30047
   150
wenzelm@30047
   151
text {*
wenzelm@30047
   152
  \medskip\noindent This Isar reasoning pattern again refers to the
wenzelm@30047
   153
  primitive rule depicted above.  The system determines it in the
wenzelm@30063
   154
  ``@{command proof}'' step, which could have been spelt out more
wenzelm@30063
   155
  explicitly as ``@{command proof}~@{text "(rule InterI)"}''.  Note
wenzelm@30061
   156
  that the rule involves both a local parameter @{term "A"} and an
wenzelm@30055
   157
  assumption @{prop "A \<in> \<A>"} in the nested reasoning.  This kind of
wenzelm@30055
   158
  compound rule typically demands a genuine sub-proof in Isar, working
wenzelm@30055
   159
  backwards rather than forwards as seen before.  In the proof body we
wenzelm@30063
   160
  encounter the @{command fix}-@{command assume}-@{command show}
wenzelm@30063
   161
  outline of nested sub-proofs that is typical for Isar.  The final
wenzelm@30063
   162
  @{command show} is like @{command have} followed by an additional
wenzelm@30063
   163
  refinement of the enclosing claim, using the rule derived from the
wenzelm@30063
   164
  proof body.
wenzelm@30047
   165
wenzelm@30047
   166
  \medskip The next example involves @{term "\<Union>\<A>"}, which can be
wenzelm@30047
   167
  characterized as the set of all @{term "x"} such that @{prop "\<exists>A. x
wenzelm@30047
   168
  \<in> A \<and> A \<in> \<A>"}.  The elimination rule for @{prop "x \<in> \<Union>\<A>"} does
wenzelm@30047
   169
  not mention @{text "\<exists>"} and @{text "\<and>"} at all, but admits to obtain
wenzelm@30047
   170
  directly a local @{term "A"} such that @{prop "x \<in> A"} and @{prop "A
wenzelm@30047
   171
  \<in> \<A>"} hold.  This corresponds to the following Isar proof and
wenzelm@30047
   172
  inference rule, respectively:
wenzelm@30047
   173
*}
wenzelm@30047
   174
wenzelm@30047
   175
text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
wenzelm@30047
   176
wenzelm@30047
   177
(*<*)
wenzelm@36367
   178
example_proof
wenzelm@30047
   179
(*>*)
wenzelm@30047
   180
    assume "x \<in> \<Union>\<A>"
wenzelm@30047
   181
    then have C
wenzelm@30047
   182
    proof
wenzelm@30047
   183
      fix A
wenzelm@30047
   184
      assume "x \<in> A" and "A \<in> \<A>"
wenzelm@30059
   185
      show C sorry %noproof
wenzelm@30047
   186
    qed
wenzelm@30047
   187
(*<*)
wenzelm@30047
   188
qed
wenzelm@30047
   189
(*>*)
wenzelm@30047
   190
wenzelm@30047
   191
text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
wenzelm@30047
   192
wenzelm@30047
   193
text {*
wenzelm@30047
   194
  \infer{@{prop "C"}}{@{prop "x \<in> \<Union>\<A>"} & \infer*{@{prop "C"}~}{@{text "[A][x \<in> A, A \<in> \<A>]"}}}
wenzelm@30047
   195
*}
wenzelm@30047
   196
wenzelm@30047
   197
text_raw {*\end{minipage}*}
wenzelm@30047
   198
wenzelm@30047
   199
text {*
wenzelm@30047
   200
  \medskip\noindent Although the Isar proof follows the natural
wenzelm@30047
   201
  deduction rule closely, the text reads not as natural as
wenzelm@30047
   202
  anticipated.  There is a double occurrence of an arbitrary
wenzelm@30047
   203
  conclusion @{prop "C"}, which represents the final result, but is
wenzelm@30047
   204
  irrelevant for now.  This issue arises for any elimination rule
wenzelm@30047
   205
  involving local parameters.  Isar provides the derived language
wenzelm@30063
   206
  element @{command obtain}, which is able to perform the same
wenzelm@30047
   207
  elimination proof more conveniently:
wenzelm@30047
   208
*}
wenzelm@30047
   209
wenzelm@30047
   210
(*<*)
wenzelm@36367
   211
example_proof
wenzelm@30047
   212
(*>*)
wenzelm@30047
   213
    assume "x \<in> \<Union>\<A>"
wenzelm@30047
   214
    then obtain A where "x \<in> A" and "A \<in> \<A>" ..
wenzelm@30047
   215
(*<*)
wenzelm@30047
   216
qed
wenzelm@30047
   217
(*>*)
wenzelm@30047
   218
wenzelm@30047
   219
text {*
wenzelm@30047
   220
  \noindent Here we avoid to mention the final conclusion @{prop "C"}
wenzelm@30047
   221
  and return to plain forward reasoning.  The rule involved in the
wenzelm@30047
   222
  ``@{command ".."}'' proof is the same as before.
wenzelm@30042
   223
*}
wenzelm@30042
   224
wenzelm@30055
   225
wenzelm@30055
   226
section {* The Pure framework \label{sec:framework-pure} *}
wenzelm@30055
   227
wenzelm@30055
   228
text {*
wenzelm@30055
   229
  The Pure logic \cite{paulson-found,paulson700} is an intuitionistic
wenzelm@30055
   230
  fragment of higher-order logic \cite{church40}.  In type-theoretic
wenzelm@30055
   231
  parlance, there are three levels of @{text "\<lambda>"}-calculus with
wenzelm@30061
   232
  corresponding arrows @{text "\<Rightarrow>"}/@{text "\<And>"}/@{text "\<Longrightarrow>"}:
wenzelm@30055
   233
wenzelm@30061
   234
  \medskip
wenzelm@30061
   235
  \begin{tabular}{ll}
wenzelm@30061
   236
  @{text "\<alpha> \<Rightarrow> \<beta>"} & syntactic function space (terms depending on terms) \\
wenzelm@30061
   237
  @{text "\<And>x. B(x)"} & universal quantification (proofs depending on terms) \\
wenzelm@30061
   238
  @{text "A \<Longrightarrow> B"} & implication (proofs depending on proofs) \\
wenzelm@30061
   239
  \end{tabular}
wenzelm@30061
   240
  \medskip
wenzelm@30061
   241
wenzelm@30061
   242
  \noindent Here only the types of syntactic terms, and the
wenzelm@30061
   243
  propositions of proof terms have been shown.  The @{text
wenzelm@30061
   244
  "\<lambda>"}-structure of proofs can be recorded as an optional feature of
wenzelm@30061
   245
  the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but
wenzelm@30061
   246
  the formal system can never depend on them due to \emph{proof
wenzelm@30061
   247
  irrelevance}.
wenzelm@30061
   248
wenzelm@30061
   249
  On top of this most primitive layer of proofs, Pure implements a
wenzelm@30061
   250
  generic calculus for nested natural deduction rules, similar to
wenzelm@30061
   251
  \cite{Schroeder-Heister:1984}.  Here object-logic inferences are
wenzelm@30061
   252
  internalized as formulae over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
wenzelm@30061
   253
  Combining such rule statements may involve higher-order unification
wenzelm@30061
   254
  \cite{paulson-natural}.
wenzelm@30055
   255
*}
wenzelm@30055
   256
wenzelm@30055
   257
wenzelm@30055
   258
subsection {* Primitive inferences *}
wenzelm@30055
   259
wenzelm@30055
   260
text {*
wenzelm@30055
   261
  Term syntax provides explicit notation for abstraction @{text "\<lambda>x ::
wenzelm@30055
   262
  \<alpha>. b(x)"} and application @{text "b a"}, while types are usually
wenzelm@30055
   263
  implicit thanks to type-inference; terms of type @{text "prop"} are
wenzelm@30055
   264
  called propositions.  Logical statements are composed via @{text "\<And>x
wenzelm@30055
   265
  :: \<alpha>. B(x)"} and @{text "A \<Longrightarrow> B"}.  Primitive reasoning operates on
wenzelm@30055
   266
  judgments of the form @{text "\<Gamma> \<turnstile> \<phi>"}, with standard introduction
wenzelm@30055
   267
  and elimination rules for @{text "\<And>"} and @{text "\<Longrightarrow>"} that refer to
wenzelm@30055
   268
  fixed parameters @{text "x\<^isub>1, \<dots>, x\<^isub>m"} and hypotheses
wenzelm@30055
   269
  @{text "A\<^isub>1, \<dots>, A\<^isub>n"} from the context @{text "\<Gamma>"};
wenzelm@30055
   270
  the corresponding proof terms are left implicit.  The subsequent
wenzelm@30055
   271
  inference rules define @{text "\<Gamma> \<turnstile> \<phi>"} inductively, relative to a
wenzelm@30055
   272
  collection of axioms:
wenzelm@30055
   273
wenzelm@30055
   274
  \[
wenzelm@30055
   275
  \infer{@{text "\<turnstile> A"}}{(@{text "A"} \text{~axiom})}
wenzelm@30055
   276
  \qquad
wenzelm@30055
   277
  \infer{@{text "A \<turnstile> A"}}{}
wenzelm@30055
   278
  \]
wenzelm@30055
   279
wenzelm@30055
   280
  \[
wenzelm@30055
   281
  \infer{@{text "\<Gamma> \<turnstile> \<And>x. B(x)"}}{@{text "\<Gamma> \<turnstile> B(x)"} & @{text "x \<notin> \<Gamma>"}}
wenzelm@30055
   282
  \qquad
wenzelm@30055
   283
  \infer{@{text "\<Gamma> \<turnstile> B(a)"}}{@{text "\<Gamma> \<turnstile> \<And>x. B(x)"}}
wenzelm@30055
   284
  \]
wenzelm@30055
   285
wenzelm@30055
   286
  \[
wenzelm@30055
   287
  \infer{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
wenzelm@30055
   288
  \qquad
wenzelm@30055
   289
  \infer{@{text "\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B"}}{@{text "\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B"} & @{text "\<Gamma>\<^sub>2 \<turnstile> A"}}
wenzelm@30055
   290
  \]
wenzelm@30055
   291
wenzelm@30055
   292
  Furthermore, Pure provides a built-in equality @{text "\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow>
wenzelm@30055
   293
  prop"} with axioms for reflexivity, substitution, extensionality,
wenzelm@30055
   294
  and @{text "\<alpha>\<beta>\<eta>"}-conversion on @{text "\<lambda>"}-terms.
wenzelm@30055
   295
wenzelm@30055
   296
  \medskip An object-logic introduces another layer on top of Pure,
wenzelm@30055
   297
  e.g.\ with types @{text "i"} for individuals and @{text "o"} for
wenzelm@30055
   298
  propositions, term constants @{text "Trueprop :: o \<Rightarrow> prop"} as
wenzelm@30055
   299
  (implicit) derivability judgment and connectives like @{text "\<and> :: o
wenzelm@30055
   300
  \<Rightarrow> o \<Rightarrow> o"} or @{text "\<forall> :: (i \<Rightarrow> o) \<Rightarrow> o"}, and axioms for object-level
wenzelm@30055
   301
  rules such as @{text "conjI: A \<Longrightarrow> B \<Longrightarrow> A \<and> B"} or @{text "allI: (\<And>x. B
wenzelm@30055
   302
  x) \<Longrightarrow> \<forall>x. B x"}.  Derived object rules are represented as theorems of
wenzelm@30055
   303
  Pure.  After the initial object-logic setup, further axiomatizations
wenzelm@30055
   304
  are usually avoided; plain definitions and derived principles are
wenzelm@30055
   305
  used exclusively.
wenzelm@30055
   306
*}
wenzelm@30055
   307
wenzelm@30055
   308
wenzelm@30055
   309
subsection {* Reasoning with rules \label{sec:framework-resolution} *}
wenzelm@30055
   310
wenzelm@30055
   311
text {*
wenzelm@30055
   312
  Primitive inferences mostly serve foundational purposes.  The main
wenzelm@30055
   313
  reasoning mechanisms of Pure operate on nested natural deduction
wenzelm@30055
   314
  rules expressed as formulae, using @{text "\<And>"} to bind local
wenzelm@30055
   315
  parameters and @{text "\<Longrightarrow>"} to express entailment.  Multiple
wenzelm@30055
   316
  parameters and premises are represented by repeating these
wenzelm@30061
   317
  connectives in a right-associative manner.
wenzelm@30055
   318
wenzelm@30055
   319
  Since @{text "\<And>"} and @{text "\<Longrightarrow>"} commute thanks to the theorem
wenzelm@30055
   320
  @{prop "(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)"}, we may assume w.l.o.g.\
wenzelm@30055
   321
  that rule statements always observe the normal form where
wenzelm@30055
   322
  quantifiers are pulled in front of implications at each level of
wenzelm@30055
   323
  nesting.  This means that any Pure proposition may be presented as a
wenzelm@30055
   324
  \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the
wenzelm@30055
   325
  form @{text "\<And>x\<^isub>1 \<dots> x\<^isub>m. H\<^isub>1 \<Longrightarrow> \<dots> H\<^isub>n \<Longrightarrow>
wenzelm@30061
   326
  A"} for @{text "m, n \<ge> 0"}, and @{text "A"} atomic, and @{text
wenzelm@30061
   327
  "H\<^isub>1, \<dots>, H\<^isub>n"} being recursively of the same format.
wenzelm@30055
   328
  Following the convention that outermost quantifiers are implicit,
wenzelm@30055
   329
  Horn clauses @{text "A\<^isub>1 \<Longrightarrow> \<dots> A\<^isub>n \<Longrightarrow> A"} are a special
wenzelm@30055
   330
  case of this.
wenzelm@30055
   331
wenzelm@30061
   332
  For example, @{text "\<inter>"}-introduction rule encountered before is
wenzelm@30061
   333
  represented as a Pure theorem as follows:
wenzelm@30061
   334
  \[
wenzelm@30061
   335
  @{text "IntI:"}~@{prop "x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A \<inter> B"}
wenzelm@30061
   336
  \]
wenzelm@30061
   337
wenzelm@30061
   338
  \noindent This is a plain Horn clause, since no further nesting on
wenzelm@30061
   339
  the left is involved.  The general @{text "\<Inter>"}-introduction
wenzelm@30061
   340
  corresponds to a Hereditary Harrop Formula with one additional level
wenzelm@30061
   341
  of nesting:
wenzelm@30061
   342
  \[
wenzelm@30061
   343
  @{text "InterI:"}~@{prop "(\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A) \<Longrightarrow> x \<in> \<Inter>\<A>"}
wenzelm@30061
   344
  \]
wenzelm@30061
   345
wenzelm@30055
   346
  \medskip Goals are also represented as rules: @{text "A\<^isub>1 \<Longrightarrow>
wenzelm@30055
   347
  \<dots> A\<^isub>n \<Longrightarrow> C"} states that the sub-goals @{text "A\<^isub>1, \<dots>,
wenzelm@30055
   348
  A\<^isub>n"} entail the result @{text "C"}; for @{text "n = 0"} the
wenzelm@30055
   349
  goal is finished.  To allow @{text "C"} being a rule statement
wenzelm@30055
   350
  itself, we introduce the protective marker @{text "# :: prop \<Rightarrow>
wenzelm@30055
   351
  prop"}, which is defined as identity and hidden from the user.  We
wenzelm@30055
   352
  initialize and finish goal states as follows:
wenzelm@30055
   353
wenzelm@30055
   354
  \[
wenzelm@30055
   355
  \begin{array}{c@ {\qquad}c}
wenzelm@30055
   356
  \infer[(@{inference_def init})]{@{text "C \<Longrightarrow> #C"}}{} &
wenzelm@30055
   357
  \infer[(@{inference_def finish})]{@{text C}}{@{text "#C"}}
wenzelm@30055
   358
  \end{array}
wenzelm@30055
   359
  \]
wenzelm@30055
   360
wenzelm@30061
   361
  \noindent Goal states are refined in intermediate proof steps until
wenzelm@30061
   362
  a finished form is achieved.  Here the two main reasoning principles
wenzelm@30061
   363
  are @{inference resolution}, for back-chaining a rule against a
wenzelm@30061
   364
  sub-goal (replacing it by zero or more sub-goals), and @{inference
wenzelm@30055
   365
  assumption}, for solving a sub-goal (finding a short-circuit with
wenzelm@30055
   366
  local assumptions).  Below @{text "\<^vec>x"} stands for @{text
wenzelm@30055
   367
  "x\<^isub>1, \<dots>, x\<^isub>n"} (@{text "n \<ge> 0"}).
wenzelm@30055
   368
wenzelm@30055
   369
  \[
wenzelm@30055
   370
  \infer[(@{inference_def resolution})]
wenzelm@30055
   371
  {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>A (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
wenzelm@30055
   372
  {\begin{tabular}{rl}
wenzelm@30055
   373
    @{text "rule:"} &
wenzelm@30055
   374
    @{text "\<^vec>A \<^vec>a \<Longrightarrow> B \<^vec>a"} \\
wenzelm@30055
   375
    @{text "goal:"} &
wenzelm@30055
   376
    @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
wenzelm@30055
   377
    @{text "goal unifier:"} &
wenzelm@30055
   378
    @{text "(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
wenzelm@30055
   379
   \end{tabular}}
wenzelm@30055
   380
  \]
wenzelm@30055
   381
wenzelm@30055
   382
  \medskip
wenzelm@30055
   383
wenzelm@30055
   384
  \[
wenzelm@30055
   385
  \infer[(@{inference_def assumption})]{@{text "C\<vartheta>"}}
wenzelm@30055
   386
  {\begin{tabular}{rl}
wenzelm@30055
   387
    @{text "goal:"} &
wenzelm@30055
   388
    @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C"} \\
wenzelm@30055
   389
    @{text "assm unifier:"} & @{text "A\<vartheta> = H\<^sub>i\<vartheta>"}~~\text{(for some~@{text "H\<^sub>i"})} \\
wenzelm@30055
   390
   \end{tabular}}
wenzelm@30055
   391
  \]
wenzelm@30055
   392
wenzelm@30055
   393
  The following trace illustrates goal-oriented reasoning in
wenzelm@30055
   394
  Isabelle/Pure:
wenzelm@30055
   395
wenzelm@30061
   396
  {\footnotesize
wenzelm@30055
   397
  \medskip
wenzelm@30061
   398
  \begin{tabular}{r@ {\quad}l}
wenzelm@30055
   399
  @{text "(A \<and> B \<Longrightarrow> B \<and> A) \<Longrightarrow> #(A \<and> B \<Longrightarrow> B \<and> A)"} & @{text "(init)"} \\
wenzelm@30055
   400
  @{text "(A \<and> B \<Longrightarrow> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution B \<Longrightarrow> A \<Longrightarrow> B \<and> A)"} \\
wenzelm@30055
   401
  @{text "(A \<and> B \<Longrightarrow> A \<and> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution A \<and> B \<Longrightarrow> B)"} \\
wenzelm@30055
   402
  @{text "(A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(assumption)"} \\
wenzelm@30055
   403
  @{text "(A \<and> B \<Longrightarrow> B \<and> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution A \<and> B \<Longrightarrow> A)"} \\
wenzelm@30055
   404
  @{text "#\<dots>"} & @{text "(assumption)"} \\
wenzelm@30055
   405
  @{text "A \<and> B \<Longrightarrow> B \<and> A"} & @{text "(finish)"} \\
wenzelm@30055
   406
  \end{tabular}
wenzelm@30055
   407
  \medskip
wenzelm@30061
   408
  }
wenzelm@30055
   409
wenzelm@30055
   410
  Compositions of @{inference assumption} after @{inference
wenzelm@30055
   411
  resolution} occurs quite often, typically in elimination steps.
wenzelm@30055
   412
  Traditional Isabelle tactics accommodate this by a combined
wenzelm@30055
   413
  @{inference_def elim_resolution} principle.  In contrast, Isar uses
wenzelm@30055
   414
  a slightly more refined combination, where the assumptions to be
wenzelm@30055
   415
  closed are marked explicitly, using again the protective marker
wenzelm@30055
   416
  @{text "#"}:
wenzelm@30055
   417
wenzelm@30055
   418
  \[
wenzelm@30055
   419
  \infer[(@{inference refinement})]
wenzelm@30055
   420
  {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>G' (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
wenzelm@30055
   421
  {\begin{tabular}{rl}
wenzelm@30055
   422
    @{text "sub\<dash>proof:"} &
wenzelm@30055
   423
    @{text "\<^vec>G \<^vec>a \<Longrightarrow> B \<^vec>a"} \\
wenzelm@30055
   424
    @{text "goal:"} &
wenzelm@30055
   425
    @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
wenzelm@30055
   426
    @{text "goal unifier:"} &
wenzelm@30055
   427
    @{text "(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
wenzelm@30055
   428
    @{text "assm unifiers:"} &
wenzelm@30055
   429
    @{text "(\<lambda>\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\<vartheta> = #H\<^sub>i\<vartheta>"} \\
wenzelm@30055
   430
    & \quad (for each marked @{text "G\<^sub>j"} some @{text "#H\<^sub>i"}) \\
wenzelm@30055
   431
   \end{tabular}}
wenzelm@30055
   432
  \]
wenzelm@30055
   433
wenzelm@30055
   434
  \noindent Here the @{text "sub\<dash>proof"} rule stems from the
wenzelm@30063
   435
  main @{command fix}-@{command assume}-@{command show} outline of
wenzelm@30063
   436
  Isar (cf.\ \secref{sec:framework-subproof}): each assumption
wenzelm@30055
   437
  indicated in the text results in a marked premise @{text "G"} above.
wenzelm@30061
   438
  The marking enforces resolution against one of the sub-goal's
wenzelm@30063
   439
  premises.  Consequently, @{command fix}-@{command assume}-@{command
wenzelm@30063
   440
  show} enables to fit the result of a sub-proof quite robustly into a
wenzelm@30063
   441
  pending sub-goal, while maintaining a good measure of flexibility.
wenzelm@30055
   442
*}
wenzelm@30055
   443
wenzelm@30055
   444
wenzelm@30055
   445
section {* The Isar proof language \label{sec:framework-isar} *}
wenzelm@30055
   446
wenzelm@30055
   447
text {*
wenzelm@30055
   448
  Structured proofs are presented as high-level expressions for
wenzelm@30055
   449
  composing entities of Pure (propositions, facts, and goals).  The
wenzelm@30055
   450
  Isar proof language allows to organize reasoning within the
wenzelm@30055
   451
  underlying rule calculus of Pure, but Isar is not another logical
wenzelm@30055
   452
  calculus!
wenzelm@30055
   453
wenzelm@30055
   454
  Isar is an exercise in sound minimalism.  Approximately half of the
wenzelm@30055
   455
  language is introduced as primitive, the rest defined as derived
wenzelm@30055
   456
  concepts.  The following grammar describes the core language
wenzelm@30055
   457
  (category @{text "proof"}), which is embedded into theory
wenzelm@30055
   458
  specification elements such as @{command theorem}; see also
wenzelm@30055
   459
  \secref{sec:framework-stmt} for the separate category @{text
wenzelm@30055
   460
  "statement"}.
wenzelm@30055
   461
wenzelm@30055
   462
  \medskip
wenzelm@30055
   463
  \begin{tabular}{rcl}
wenzelm@30055
   464
    @{text "theory\<dash>stmt"} & = & @{command "theorem"}~@{text "statement proof  |"}~~@{command "definition"}~@{text "\<dots>  |  \<dots>"} \\[1ex]
wenzelm@30055
   465
wenzelm@30055
   466
    @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\[1ex]
wenzelm@30055
   467
wenzelm@30055
   468
    @{text prfx} & = & @{command "using"}~@{text "facts"} \\
wenzelm@30055
   469
    & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\
wenzelm@30055
   470
wenzelm@30055
   471
    @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\
wenzelm@30055
   472
    & @{text "|"} & @{command "next"} \\
wenzelm@30055
   473
    & @{text "|"} & @{command "note"}~@{text "name = facts"} \\
wenzelm@30055
   474
    & @{text "|"} & @{command "let"}~@{text "term = term"} \\
wenzelm@30055
   475
    & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\
wenzelm@30063
   476
    & @{text "|"} & @{command assume}~@{text "\<guillemotleft>inference\<guillemotright> name: props"} \\
wenzelm@30055
   477
    & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\
wenzelm@30055
   478
    @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\
wenzelm@30055
   479
    & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\
wenzelm@30055
   480
  \end{tabular}
wenzelm@30055
   481
wenzelm@30055
   482
  \medskip Simultaneous propositions or facts may be separated by the
wenzelm@30055
   483
  @{keyword "and"} keyword.
wenzelm@30055
   484
wenzelm@30055
   485
  \medskip The syntax for terms and propositions is inherited from
wenzelm@30055
   486
  Pure (and the object-logic).  A @{text "pattern"} is a @{text
wenzelm@30055
   487
  "term"} with schematic variables, to be bound by higher-order
wenzelm@30055
   488
  matching.
wenzelm@30055
   489
wenzelm@30061
   490
  \medskip Facts may be referenced by name or proposition.  For
wenzelm@30063
   491
  example, the result of ``@{command have}~@{text "a: A \<langle>proof\<rangle>"}''
wenzelm@30061
   492
  becomes available both as @{text "a"} and
wenzelm@30061
   493
  \isacharbackquoteopen@{text "A"}\isacharbackquoteclose.  Moreover,
wenzelm@30061
   494
  fact expressions may involve attributes that modify either the
wenzelm@30061
   495
  theorem or the background context.  For example, the expression
wenzelm@30061
   496
  ``@{text "a [OF b]"}'' refers to the composition of two facts
wenzelm@30061
   497
  according to the @{inference resolution} inference of
wenzelm@30061
   498
  \secref{sec:framework-resolution}, while ``@{text "a [intro]"}''
wenzelm@30061
   499
  declares a fact as introduction rule in the context.
wenzelm@30055
   500
wenzelm@30061
   501
  The special fact called ``@{fact this}'' always refers to the last
wenzelm@30063
   502
  result, as produced by @{command note}, @{command assume}, @{command
wenzelm@30063
   503
  have}, or @{command show}.  Since @{command note} occurs
wenzelm@30063
   504
  frequently together with @{command then} we provide some
wenzelm@30063
   505
  abbreviations:
wenzelm@30055
   506
wenzelm@30063
   507
  \medskip
wenzelm@30063
   508
  \begin{tabular}{rcl}
wenzelm@30063
   509
    @{command from}~@{text a} & @{text "\<equiv>"} & @{command note}~@{text a}~@{command then} \\
wenzelm@30063
   510
    @{command with}~@{text a} & @{text "\<equiv>"} & @{command from}~@{text "a \<AND> this"} \\
wenzelm@30063
   511
  \end{tabular}
wenzelm@30063
   512
  \medskip
wenzelm@30063
   513
wenzelm@30063
   514
  The @{text "method"} category is essentially a parameter and may be
wenzelm@30063
   515
  populated later.  Methods use the facts indicated by @{command
wenzelm@30063
   516
  "then"} or @{command using}, and then operate on the goal state.
wenzelm@30063
   517
  Some basic methods are predefined: ``@{method "-"}'' leaves the goal
wenzelm@30063
   518
  unchanged, ``@{method this}'' applies the facts as rules to the
wenzelm@30063
   519
  goal, ``@{method "rule"}'' applies the facts to another rule and the
wenzelm@30063
   520
  result to the goal (both ``@{method this}'' and ``@{method rule}''
wenzelm@30063
   521
  refer to @{inference resolution} of
wenzelm@30055
   522
  \secref{sec:framework-resolution}).  The secondary arguments to
wenzelm@30055
   523
  ``@{method rule}'' may be specified explicitly as in ``@{text "(rule
wenzelm@30055
   524
  a)"}'', or picked from the context.  In the latter case, the system
wenzelm@30055
   525
  first tries rules declared as @{attribute (Pure) elim} or
wenzelm@30055
   526
  @{attribute (Pure) dest}, followed by those declared as @{attribute
wenzelm@30055
   527
  (Pure) intro}.
wenzelm@30055
   528
wenzelm@30063
   529
  The default method for @{command proof} is ``@{method rule}''
wenzelm@30063
   530
  (arguments picked from the context), for @{command qed} it is
wenzelm@30055
   531
  ``@{method "-"}''.  Further abbreviations for terminal proof steps
wenzelm@30055
   532
  are ``@{command "by"}~@{text "method\<^sub>1 method\<^sub>2"}'' for
wenzelm@30063
   533
  ``@{command proof}~@{text "method\<^sub>1"}~@{command qed}~@{text
wenzelm@30063
   534
  "method\<^sub>2"}'', and ``@{command ".."}'' for ``@{command
wenzelm@30063
   535
  "by"}~@{method rule}, and ``@{command "."}'' for ``@{command
wenzelm@30063
   536
  "by"}~@{method this}''.  The @{command unfolding} element operates
wenzelm@30063
   537
  directly on the current facts and goal by applying equalities.
wenzelm@30055
   538
wenzelm@30063
   539
  \medskip Block structure can be indicated explicitly by ``@{command
wenzelm@30063
   540
  "{"}~@{text "\<dots>"}~@{command "}"}'', although the body of a sub-proof
wenzelm@30063
   541
  already involves implicit nesting.  In any case, @{command next}
wenzelm@30063
   542
  jumps into the next section of a block, i.e.\ it acts like closing
wenzelm@30063
   543
  an implicit block scope and opening another one; there is no direct
wenzelm@30063
   544
  correspondence to subgoals here.
wenzelm@30055
   545
wenzelm@30063
   546
  The remaining elements @{command fix} and @{command assume} build up
wenzelm@30063
   547
  a local context (see \secref{sec:framework-context}), while
wenzelm@30063
   548
  @{command show} refines a pending sub-goal by the rule resulting
wenzelm@30055
   549
  from a nested sub-proof (see \secref{sec:framework-subproof}).
wenzelm@30055
   550
  Further derived concepts will support calculational reasoning (see
wenzelm@30055
   551
  \secref{sec:framework-calc}).
wenzelm@30055
   552
*}
wenzelm@30055
   553
wenzelm@30055
   554
wenzelm@30055
   555
subsection {* Context elements \label{sec:framework-context} *}
wenzelm@30055
   556
wenzelm@30055
   557
text {*
wenzelm@30055
   558
  In judgments @{text "\<Gamma> \<turnstile> \<phi>"} of the primitive framework, @{text "\<Gamma>"}
wenzelm@30055
   559
  essentially acts like a proof context.  Isar elaborates this idea
wenzelm@30061
   560
  towards a higher-level notion, with additional information for
wenzelm@30055
   561
  type-inference, term abbreviations, local facts, hypotheses etc.
wenzelm@30055
   562
wenzelm@30063
   563
  The element @{command fix}~@{text "x :: \<alpha>"} declares a local
wenzelm@30055
   564
  parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in
wenzelm@30055
   565
  results exported from the context, @{text "x"} may become anything.
wenzelm@30063
   566
  The @{command assume}~@{text "\<guillemotleft>inference\<guillemotright>"} element provides a
wenzelm@30063
   567
  general interface to hypotheses: ``@{command assume}~@{text
wenzelm@30063
   568
  "\<guillemotleft>inference\<guillemotright> A"}'' produces @{text "A \<turnstile> A"} locally, while the
wenzelm@30063
   569
  included inference tells how to discharge @{text A} from results
wenzelm@30063
   570
  @{text "A \<turnstile> B"} later on.  There is no user-syntax for @{text
wenzelm@30063
   571
  "\<guillemotleft>inference\<guillemotright>"}, i.e.\ it may only occur internally when derived
wenzelm@30063
   572
  commands are defined in ML.
wenzelm@30063
   573
wenzelm@30063
   574
  At the user-level, the default inference for @{command assume} is
wenzelm@30063
   575
  @{inference discharge} as given below.  The additional variants
wenzelm@30063
   576
  @{command presume} and @{command def} are defined as follows:
wenzelm@30063
   577
wenzelm@30063
   578
  \medskip
wenzelm@30063
   579
  \begin{tabular}{rcl}
wenzelm@30063
   580
    @{command presume}~@{text A} & @{text "\<equiv>"} & @{command assume}~@{text "\<guillemotleft>weak\<dash>discharge\<guillemotright> A"} \\
wenzelm@30063
   581
    @{command def}~@{text "x \<equiv> a"} & @{text "\<equiv>"} & @{command fix}~@{text x}~@{command assume}~@{text "\<guillemotleft>expansion\<guillemotright> x \<equiv> a"} \\
wenzelm@30063
   582
  \end{tabular}
wenzelm@30063
   583
  \medskip
wenzelm@30055
   584
wenzelm@30055
   585
  \[
wenzelm@30063
   586
  \infer[(@{inference_def discharge})]{@{text "\<strut>\<Gamma> - A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<strut>\<Gamma> \<turnstile> B"}}
wenzelm@30063
   587
  \]
wenzelm@30063
   588
  \[
wenzelm@30063
   589
  \infer[(@{inference_def "weak\<dash>discharge"})]{@{text "\<strut>\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<strut>\<Gamma> \<turnstile> B"}}
wenzelm@30063
   590
  \]
wenzelm@30063
   591
  \[
wenzelm@30055
   592
  \infer[(@{inference_def expansion})]{@{text "\<strut>\<Gamma> - (x \<equiv> a) \<turnstile> B a"}}{@{text "\<strut>\<Gamma> \<turnstile> B x"}}
wenzelm@30055
   593
  \]
wenzelm@30055
   594
wenzelm@30063
   595
  \medskip Note that @{inference discharge} and @{inference
wenzelm@30063
   596
  "weak\<dash>discharge"} differ in the marker for @{prop A}, which is
wenzelm@30063
   597
  relevant when the result of a @{command fix}-@{command
wenzelm@30063
   598
  assume}-@{command show} outline is composed with a pending goal,
wenzelm@30063
   599
  cf.\ \secref{sec:framework-subproof}.
wenzelm@30055
   600
wenzelm@30063
   601
  The most interesting derived context element in Isar is @{command
wenzelm@30063
   602
  obtain} \cite[\S5.3]{Wenzel-PhD}, which supports generalized
wenzelm@30063
   603
  elimination steps in a purely forward manner.  The @{command obtain}
wenzelm@30065
   604
  command takes a specification of parameters @{text "\<^vec>x"} and
wenzelm@30063
   605
  assumptions @{text "\<^vec>A"} to be added to the context, together
wenzelm@30063
   606
  with a proof of a case rule stating that this extension is
wenzelm@30063
   607
  conservative (i.e.\ may be removed from closed results later on):
wenzelm@30055
   608
wenzelm@30055
   609
  \medskip
wenzelm@30055
   610
  \begin{tabular}{l}
wenzelm@30055
   611
  @{text "\<langle>facts\<rangle>"}~~@{command obtain}~@{text "\<^vec>x \<WHERE> \<^vec>A \<^vec>x  \<langle>proof\<rangle> \<equiv>"} \\[0.5ex]
wenzelm@30055
   612
  \quad @{command have}~@{text "case: \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis\<rangle>"} \\
wenzelm@30055
   613
  \quad @{command proof}~@{method "-"} \\
wenzelm@30055
   614
  \qquad @{command fix}~@{text thesis} \\
wenzelm@30055
   615
  \qquad @{command assume}~@{text "[intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis"} \\
wenzelm@30059
   616
  \qquad @{command show}~@{text thesis}~@{command using}~@{text "\<langle>facts\<rangle> \<langle>proof\<rangle>"} \\
wenzelm@30055
   617
  \quad @{command qed} \\
wenzelm@30063
   618
  \quad @{command fix}~@{text "\<^vec>x"}~@{command assume}~@{text "\<guillemotleft>elimination case\<guillemotright> \<^vec>A \<^vec>x"} \\
wenzelm@30055
   619
  \end{tabular}
wenzelm@30055
   620
  \medskip
wenzelm@30055
   621
wenzelm@30055
   622
  \[
wenzelm@30055
   623
  \infer[(@{inference elimination})]{@{text "\<Gamma> \<turnstile> B"}}{
wenzelm@30055
   624
    \begin{tabular}{rl}
wenzelm@30055
   625
    @{text "case:"} &
wenzelm@30055
   626
    @{text "\<Gamma> \<turnstile> \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis"} \\[0.2ex]
wenzelm@30055
   627
    @{text "result:"} &
wenzelm@30055
   628
    @{text "\<Gamma> \<union> \<^vec>A \<^vec>y \<turnstile> B"} \\[0.2ex]
wenzelm@30055
   629
    \end{tabular}}
wenzelm@30055
   630
  \]
wenzelm@30055
   631
wenzelm@30055
   632
  \noindent Here the name ``@{text thesis}'' is a specific convention
wenzelm@30055
   633
  for an arbitrary-but-fixed proposition; in the primitive natural
wenzelm@30055
   634
  deduction rules shown before we have occasionally used @{text C}.
wenzelm@30063
   635
  The whole statement of ``@{command obtain}~@{text x}~@{keyword
wenzelm@30055
   636
  "where"}~@{text "A x"}'' may be read as a claim that @{text "A x"}
wenzelm@30055
   637
  may be assumed for some arbitrary-but-fixed @{text "x"}.  Also note
wenzelm@30063
   638
  that ``@{command obtain}~@{text "A \<AND> B"}'' without parameters
wenzelm@30063
   639
  is similar to ``@{command have}~@{text "A \<AND> B"}'', but the
wenzelm@30063
   640
  latter involves multiple sub-goals.
wenzelm@30055
   641
wenzelm@30055
   642
  \medskip The subsequent Isar proof texts explain all context
wenzelm@30055
   643
  elements introduced above using the formal proof language itself.
wenzelm@30055
   644
  After finishing a local proof within a block, we indicate the
wenzelm@30065
   645
  exported result via @{command note}.
wenzelm@30055
   646
*}
wenzelm@30055
   647
wenzelm@30055
   648
(*<*)
wenzelm@30055
   649
theorem True
wenzelm@30055
   650
proof
wenzelm@30055
   651
(*>*)
wenzelm@30061
   652
  txt_raw {* \begin{minipage}[t]{0.4\textwidth} *}
wenzelm@30055
   653
  {
wenzelm@30055
   654
    fix x
wenzelm@30061
   655
    have "B x" sorry %noproof
wenzelm@30055
   656
  }
wenzelm@30055
   657
  note `\<And>x. B x`
wenzelm@30061
   658
  txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.4\textwidth} *}(*<*)next(*>*)
wenzelm@30061
   659
  {
wenzelm@30061
   660
    assume A
wenzelm@30061
   661
    have B sorry %noproof
wenzelm@30061
   662
  }
wenzelm@30061
   663
  note `A \<Longrightarrow> B`
wenzelm@30061
   664
  txt_raw {* \end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} *}(*<*)next(*>*)
wenzelm@30055
   665
  {
wenzelm@30055
   666
    def x \<equiv> a
wenzelm@30061
   667
    have "B x" sorry %noproof
wenzelm@30055
   668
  }
wenzelm@30055
   669
  note `B a`
wenzelm@30061
   670
  txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.4\textwidth} *}(*<*)next(*>*)
wenzelm@30055
   671
  {
wenzelm@30061
   672
    obtain x where "A x" sorry %noproof
wenzelm@30059
   673
    have B sorry %noproof
wenzelm@30055
   674
  }
wenzelm@30055
   675
  note `B`
wenzelm@30055
   676
  txt_raw {* \end{minipage} *}
wenzelm@30055
   677
(*<*)
wenzelm@30055
   678
qed
wenzelm@30055
   679
(*>*)
wenzelm@30055
   680
wenzelm@30065
   681
text {*
wenzelm@30065
   682
  \bigskip\noindent This illustrates the meaning of Isar context
wenzelm@30065
   683
  elements without goals getting in between.
wenzelm@30065
   684
*}
wenzelm@30055
   685
wenzelm@30055
   686
subsection {* Structured statements \label{sec:framework-stmt} *}
wenzelm@30055
   687
wenzelm@30055
   688
text {*
wenzelm@30055
   689
  The category @{text "statement"} of top-level theorem specifications
wenzelm@30055
   690
  is defined as follows:
wenzelm@30055
   691
wenzelm@30055
   692
  \medskip
wenzelm@30055
   693
  \begin{tabular}{rcl}
wenzelm@30055
   694
  @{text "statement"} & @{text "\<equiv>"} & @{text "name: props \<AND> \<dots>"} \\
wenzelm@30055
   695
  & @{text "|"} & @{text "context\<^sup>* conclusion"} \\[0.5ex]
wenzelm@30055
   696
wenzelm@30055
   697
  @{text "context"} & @{text "\<equiv>"} & @{text "\<FIXES> vars \<AND> \<dots>"} \\
wenzelm@30055
   698
  & @{text "|"} & @{text "\<ASSUMES> name: props \<AND> \<dots>"} \\
wenzelm@30055
   699
wenzelm@30055
   700
  @{text "conclusion"} & @{text "\<equiv>"} & @{text "\<SHOWS> name: props \<AND> \<dots>"} \\
wenzelm@30061
   701
  & @{text "|"} & @{text "\<OBTAINS> vars \<AND> \<dots> \<WHERE> name: props \<AND> \<dots>"} \\
wenzelm@30061
   702
  & & \quad @{text "\<BBAR> \<dots>"} \\
wenzelm@30055
   703
  \end{tabular}
wenzelm@30055
   704
wenzelm@30055
   705
  \medskip\noindent A simple @{text "statement"} consists of named
wenzelm@30055
   706
  propositions.  The full form admits local context elements followed
wenzelm@30055
   707
  by the actual conclusions, such as ``@{keyword "fixes"}~@{text
wenzelm@30055
   708
  x}~@{keyword "assumes"}~@{text "A x"}~@{keyword "shows"}~@{text "B
wenzelm@30055
   709
  x"}''.  The final result emerges as a Pure rule after discharging
wenzelm@30055
   710
  the context: @{prop "\<And>x. A x \<Longrightarrow> B x"}.
wenzelm@30055
   711
wenzelm@30055
   712
  The @{keyword "obtains"} variant is another abbreviation defined
wenzelm@30055
   713
  below; unlike @{command obtain} (cf.\
wenzelm@30055
   714
  \secref{sec:framework-context}) there may be several ``cases''
wenzelm@30055
   715
  separated by ``@{text "\<BBAR>"}'', each consisting of several
wenzelm@30055
   716
  parameters (@{text "vars"}) and several premises (@{text "props"}).
wenzelm@30055
   717
  This specifies multi-branch elimination rules.
wenzelm@30055
   718
wenzelm@30055
   719
  \medskip
wenzelm@30055
   720
  \begin{tabular}{l}
wenzelm@30055
   721
  @{text "\<OBTAINS> \<^vec>x \<WHERE> \<^vec>A \<^vec>x   \<BBAR>   \<dots>   \<equiv>"} \\[0.5ex]
wenzelm@30055
   722
  \quad @{text "\<FIXES> thesis"} \\
wenzelm@30055
   723
  \quad @{text "\<ASSUMES> [intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis  \<AND>  \<dots>"} \\
wenzelm@30055
   724
  \quad @{text "\<SHOWS> thesis"} \\
wenzelm@30055
   725
  \end{tabular}
wenzelm@30055
   726
  \medskip
wenzelm@30055
   727
wenzelm@30055
   728
  Presenting structured statements in such an ``open'' format usually
wenzelm@30055
   729
  simplifies the subsequent proof, because the outer structure of the
wenzelm@30055
   730
  problem is already laid out directly.  E.g.\ consider the following
wenzelm@30055
   731
  canonical patterns for @{text "\<SHOWS>"} and @{text "\<OBTAINS>"},
wenzelm@30055
   732
  respectively:
wenzelm@30055
   733
*}
wenzelm@30055
   734
wenzelm@30055
   735
text_raw {*\begin{minipage}{0.5\textwidth}*}
wenzelm@30055
   736
wenzelm@30055
   737
theorem
wenzelm@30055
   738
  fixes x and y
wenzelm@30055
   739
  assumes "A x" and "B y"
wenzelm@30055
   740
  shows "C x y"
wenzelm@30055
   741
proof -
wenzelm@30055
   742
  from `A x` and `B y`
wenzelm@30059
   743
  show "C x y" sorry %noproof
wenzelm@30055
   744
qed
wenzelm@30055
   745
wenzelm@30055
   746
text_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
wenzelm@30055
   747
wenzelm@30055
   748
theorem
wenzelm@30055
   749
  obtains x and y
wenzelm@30055
   750
  where "A x" and "B y"
wenzelm@30055
   751
proof -
wenzelm@30059
   752
  have "A a" and "B b" sorry %noproof
wenzelm@30055
   753
  then show thesis ..
wenzelm@30055
   754
qed
wenzelm@30055
   755
wenzelm@30055
   756
text_raw {*\end{minipage}*}
wenzelm@30055
   757
wenzelm@30055
   758
text {*
wenzelm@30055
   759
  \medskip\noindent Here local facts \isacharbackquoteopen@{text "A
wenzelm@30055
   760
  x"}\isacharbackquoteclose\ and \isacharbackquoteopen@{text "B
wenzelm@30055
   761
  y"}\isacharbackquoteclose\ are referenced immediately; there is no
wenzelm@30055
   762
  need to decompose the logical rule structure again.  In the second
wenzelm@30055
   763
  proof the final ``@{command then}~@{command show}~@{text
wenzelm@30055
   764
  thesis}~@{command ".."}''  involves the local rule case @{text "\<And>x
wenzelm@30055
   765
  y. A x \<Longrightarrow> B y \<Longrightarrow> thesis"} for the particular instance of terms @{text
wenzelm@30055
   766
  "a"} and @{text "b"} produced in the body.
wenzelm@30055
   767
*}
wenzelm@30055
   768
wenzelm@30055
   769
wenzelm@30055
   770
subsection {* Structured proof refinement \label{sec:framework-subproof} *}
wenzelm@30055
   771
wenzelm@30055
   772
text {*
wenzelm@30055
   773
  By breaking up the grammar for the Isar proof language, we may
wenzelm@30055
   774
  understand a proof text as a linear sequence of individual proof
wenzelm@30055
   775
  commands.  These are interpreted as transitions of the Isar virtual
wenzelm@30055
   776
  machine (Isar/VM), which operates on a block-structured
wenzelm@30055
   777
  configuration in single steps.  This allows users to write proof
wenzelm@30055
   778
  texts in an incremental manner, and inspect intermediate
wenzelm@30055
   779
  configurations for debugging.
wenzelm@30055
   780
wenzelm@30055
   781
  The basic idea is analogous to evaluating algebraic expressions on a
wenzelm@30055
   782
  stack machine: @{text "(a + b) \<cdot> c"} then corresponds to a sequence
wenzelm@30055
   783
  of single transitions for each symbol @{text "(, a, +, b, ), \<cdot>, c"}.
wenzelm@30055
   784
  In Isar the algebraic values are facts or goals, and the operations
wenzelm@30055
   785
  are inferences.
wenzelm@30055
   786
wenzelm@30055
   787
  \medskip The Isar/VM state maintains a stack of nodes, each node
wenzelm@30055
   788
  contains the local proof context, the linguistic mode, and a pending
wenzelm@30055
   789
  goal (optional).  The mode determines the type of transition that
wenzelm@30055
   790
  may be performed next, it essentially alternates between forward and
wenzelm@30064
   791
  backward reasoning, with an intermediate stage for chained facts
wenzelm@30064
   792
  (see \figref{fig:isar-vm}).
wenzelm@30064
   793
wenzelm@30064
   794
  \begin{figure}[htb]
wenzelm@30064
   795
  \begin{center}
wenzelm@30064
   796
  \includegraphics[width=0.8\textwidth]{Thy/document/isar-vm}
wenzelm@30064
   797
  \end{center}
wenzelm@30064
   798
  \caption{Isar/VM modes}\label{fig:isar-vm}
wenzelm@30064
   799
  \end{figure}
wenzelm@30064
   800
wenzelm@30064
   801
  For example, in @{text "state"} mode Isar acts like a mathematical
wenzelm@30064
   802
  scratch-pad, accepting declarations like @{command fix}, @{command
wenzelm@30064
   803
  assume}, and claims like @{command have}, @{command show}.  A goal
wenzelm@30064
   804
  statement changes the mode to @{text "prove"}, which means that we
wenzelm@30064
   805
  may now refine the problem via @{command unfolding} or @{command
wenzelm@30064
   806
  proof}.  Then we are again in @{text "state"} mode of a proof body,
wenzelm@30064
   807
  which may issue @{command show} statements to solve pending
wenzelm@30064
   808
  sub-goals.  A concluding @{command qed} will return to the original
wenzelm@30064
   809
  @{text "state"} mode one level upwards.  The subsequent Isar/VM
wenzelm@30064
   810
  trace indicates block structure, linguistic mode, goal state, and
wenzelm@30064
   811
  inferences:
wenzelm@30055
   812
*}
wenzelm@30055
   813
wenzelm@30061
   814
text_raw {* \begingroup\footnotesize *}
wenzelm@36367
   815
(*<*)example_proof
wenzelm@30055
   816
(*>*)
wenzelm@30061
   817
  txt_raw {* \begin{minipage}[t]{0.18\textwidth} *}
wenzelm@30055
   818
  have "A \<longrightarrow> B"
wenzelm@30055
   819
  proof
wenzelm@30055
   820
    assume A
wenzelm@30055
   821
    show B
wenzelm@30059
   822
      sorry %noproof
wenzelm@30055
   823
  qed
wenzelm@30055
   824
  txt_raw {* \end{minipage}\quad
wenzelm@30061
   825
\begin{minipage}[t]{0.06\textwidth}
wenzelm@30055
   826
@{text "begin"} \\
wenzelm@30055
   827
\\
wenzelm@30055
   828
\\
wenzelm@30055
   829
@{text "begin"} \\
wenzelm@30055
   830
@{text "end"} \\
wenzelm@30055
   831
@{text "end"} \\
wenzelm@30055
   832
\end{minipage}
wenzelm@30055
   833
\begin{minipage}[t]{0.08\textwidth}
wenzelm@30055
   834
@{text "prove"} \\
wenzelm@30055
   835
@{text "state"} \\
wenzelm@30055
   836
@{text "state"} \\
wenzelm@30055
   837
@{text "prove"} \\
wenzelm@30055
   838
@{text "state"} \\
wenzelm@30055
   839
@{text "state"} \\
wenzelm@30061
   840
\end{minipage}\begin{minipage}[t]{0.35\textwidth}
wenzelm@30055
   841
@{text "(A \<longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)"} \\
wenzelm@30055
   842
@{text "(A \<Longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)"} \\
wenzelm@30055
   843
\\
wenzelm@30055
   844
\\
wenzelm@30055
   845
@{text "#(A \<longrightarrow> B)"} \\
wenzelm@30055
   846
@{text "A \<longrightarrow> B"} \\
wenzelm@30061
   847
\end{minipage}\begin{minipage}[t]{0.4\textwidth}
wenzelm@30055
   848
@{text "(init)"} \\
wenzelm@30061
   849
@{text "(resolution impI)"} \\
wenzelm@30055
   850
\\
wenzelm@30055
   851
\\
wenzelm@30055
   852
@{text "(refinement #A \<Longrightarrow> B)"} \\
wenzelm@30055
   853
@{text "(finish)"} \\
wenzelm@30055
   854
\end{minipage} *}
wenzelm@30055
   855
(*<*)
wenzelm@30055
   856
qed
wenzelm@30055
   857
(*>*)
wenzelm@30061
   858
text_raw {* \endgroup *}
wenzelm@30055
   859
wenzelm@30055
   860
text {*
wenzelm@30061
   861
  \noindent Here the @{inference refinement} inference from
wenzelm@30055
   862
  \secref{sec:framework-resolution} mediates composition of Isar
wenzelm@30055
   863
  sub-proofs nicely.  Observe that this principle incorporates some
wenzelm@30055
   864
  degree of freedom in proof composition.  In particular, the proof
wenzelm@30055
   865
  body allows parameters and assumptions to be re-ordered, or commuted
wenzelm@30055
   866
  according to Hereditary Harrop Form.  Moreover, context elements
wenzelm@30055
   867
  that are not used in a sub-proof may be omitted altogether.  For
wenzelm@30055
   868
  example:
wenzelm@30055
   869
*}
wenzelm@30055
   870
wenzelm@30055
   871
text_raw {*\begin{minipage}{0.5\textwidth}*}
wenzelm@30055
   872
wenzelm@30055
   873
(*<*)
wenzelm@36367
   874
example_proof
wenzelm@30055
   875
(*>*)
wenzelm@30055
   876
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
wenzelm@30055
   877
  proof -
wenzelm@30055
   878
    fix x and y
wenzelm@30055
   879
    assume "A x" and "B y"
wenzelm@30059
   880
    show "C x y" sorry %noproof
wenzelm@30055
   881
  qed
wenzelm@30055
   882
wenzelm@30055
   883
txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
wenzelm@30055
   884
wenzelm@30055
   885
(*<*)
wenzelm@30055
   886
next
wenzelm@30055
   887
(*>*)
wenzelm@30055
   888
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
wenzelm@30055
   889
  proof -
wenzelm@30055
   890
    fix x assume "A x"
wenzelm@30055
   891
    fix y assume "B y"
wenzelm@30059
   892
    show "C x y" sorry %noproof
wenzelm@30055
   893
  qed
wenzelm@30055
   894
wenzelm@30061
   895
txt_raw {*\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}*}
wenzelm@30055
   896
wenzelm@30055
   897
(*<*)
wenzelm@30055
   898
next
wenzelm@30055
   899
(*>*)
wenzelm@30055
   900
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
wenzelm@30055
   901
  proof -
wenzelm@30055
   902
    fix y assume "B y"
wenzelm@30055
   903
    fix x assume "A x"
wenzelm@30055
   904
    show "C x y" sorry
wenzelm@30055
   905
  qed
wenzelm@30055
   906
wenzelm@30055
   907
txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
wenzelm@30055
   908
(*<*)
wenzelm@30055
   909
next
wenzelm@30055
   910
(*>*)
wenzelm@30055
   911
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
wenzelm@30055
   912
  proof -
wenzelm@30055
   913
    fix y assume "B y"
wenzelm@30055
   914
    fix x
wenzelm@30055
   915
    show "C x y" sorry
wenzelm@30055
   916
  qed
wenzelm@30055
   917
(*<*)
wenzelm@30055
   918
qed
wenzelm@30055
   919
(*>*)
wenzelm@30055
   920
wenzelm@30055
   921
text_raw {*\end{minipage}*}
wenzelm@30055
   922
wenzelm@30055
   923
text {*
wenzelm@30061
   924
  \medskip\noindent Such ``peephole optimizations'' of Isar texts are
wenzelm@30055
   925
  practically important to improve readability, by rearranging
wenzelm@30055
   926
  contexts elements according to the natural flow of reasoning in the
wenzelm@30055
   927
  body, while still observing the overall scoping rules.
wenzelm@30055
   928
wenzelm@30055
   929
  \medskip This illustrates the basic idea of structured proof
wenzelm@30055
   930
  processing in Isar.  The main mechanisms are based on natural
wenzelm@30055
   931
  deduction rule composition within the Pure framework.  In
wenzelm@30055
   932
  particular, there are no direct operations on goal states within the
wenzelm@30055
   933
  proof body.  Moreover, there is no hidden automated reasoning
wenzelm@30055
   934
  involved, just plain unification.
wenzelm@30055
   935
*}
wenzelm@30055
   936
wenzelm@30055
   937
wenzelm@30055
   938
subsection {* Calculational reasoning \label{sec:framework-calc} *}
wenzelm@30055
   939
wenzelm@30055
   940
text {*
wenzelm@30061
   941
  The existing Isar infrastructure is sufficiently flexible to support
wenzelm@30055
   942
  calculational reasoning (chains of transitivity steps) as derived
wenzelm@30055
   943
  concept.  The generic proof elements introduced below depend on
wenzelm@30058
   944
  rules declared as @{attribute trans} in the context.  It is left to
wenzelm@30055
   945
  the object-logic to provide a suitable rule collection for mixed
wenzelm@30058
   946
  relations of @{text "="}, @{text "<"}, @{text "\<le>"}, @{text "\<subset>"},
wenzelm@30058
   947
  @{text "\<subseteq>"} etc.  Due to the flexibility of rule composition
wenzelm@30055
   948
  (\secref{sec:framework-resolution}), substitution of equals by
wenzelm@30055
   949
  equals is covered as well, even substitution of inequalities
wenzelm@30055
   950
  involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD}
wenzelm@30055
   951
  and \cite{Bauer-Wenzel:2001}.
wenzelm@30055
   952
wenzelm@30055
   953
  The generic calculational mechanism is based on the observation that
wenzelm@30061
   954
  rules such as @{text "trans:"}~@{prop "x = y \<Longrightarrow> y = z \<Longrightarrow> x = z"}
wenzelm@30061
   955
  proceed from the premises towards the conclusion in a deterministic
wenzelm@30061
   956
  fashion.  Thus we may reason in forward mode, feeding intermediate
wenzelm@30061
   957
  results into rules selected from the context.  The course of
wenzelm@30061
   958
  reasoning is organized by maintaining a secondary fact called
wenzelm@30061
   959
  ``@{fact calculation}'', apart from the primary ``@{fact this}''
wenzelm@30061
   960
  already provided by the Isar primitives.  In the definitions below,
wenzelm@30061
   961
  @{attribute OF} refers to @{inference resolution}
wenzelm@30061
   962
  (\secref{sec:framework-resolution}) with multiple rule arguments,
wenzelm@30061
   963
  and @{text "trans"} represents to a suitable rule from the context:
wenzelm@30055
   964
wenzelm@30055
   965
  \begin{matharray}{rcl}
wenzelm@30055
   966
    @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\
wenzelm@30055
   967
    @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex]
wenzelm@30055
   968
    @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\
wenzelm@30055
   969
  \end{matharray}
wenzelm@30055
   970
wenzelm@30055
   971
  \noindent The start of a calculation is determined implicitly in the
wenzelm@30055
   972
  text: here @{command also} sets @{fact calculation} to the current
wenzelm@30055
   973
  result; any subsequent occurrence will update @{fact calculation} by
wenzelm@30055
   974
  combination with the next result and a transitivity rule.  The
wenzelm@30055
   975
  calculational sequence is concluded via @{command finally}, where
wenzelm@30055
   976
  the final result is exposed for use in a concluding claim.
wenzelm@30055
   977
wenzelm@30055
   978
  Here is a canonical proof pattern, using @{command have} to
wenzelm@30055
   979
  establish the intermediate results:
wenzelm@30055
   980
*}
wenzelm@30055
   981
wenzelm@30055
   982
(*<*)
wenzelm@36367
   983
example_proof
wenzelm@30055
   984
(*>*)
wenzelm@30055
   985
  have "a = b" sorry
wenzelm@30055
   986
  also have "\<dots> = c" sorry
wenzelm@30055
   987
  also have "\<dots> = d" sorry
wenzelm@30055
   988
  finally have "a = d" .
wenzelm@30055
   989
(*<*)
wenzelm@30055
   990
qed
wenzelm@30055
   991
(*>*)
wenzelm@30055
   992
wenzelm@30055
   993
text {*
wenzelm@30055
   994
  \noindent The term ``@{text "\<dots>"}'' above is a special abbreviation
wenzelm@30055
   995
  provided by the Isabelle/Isar syntax layer: it statically refers to
wenzelm@30055
   996
  the right-hand side argument of the previous statement given in the
wenzelm@30055
   997
  text.  Thus it happens to coincide with relevant sub-expressions in
wenzelm@30055
   998
  the calculational chain, but the exact correspondence is dependent
wenzelm@30055
   999
  on the transitivity rules being involved.
wenzelm@30055
  1000
wenzelm@30055
  1001
  \medskip Symmetry rules such as @{prop "x = y \<Longrightarrow> y = x"} are like
wenzelm@30055
  1002
  transitivities with only one premise.  Isar maintains a separate
wenzelm@30055
  1003
  rule collection declared via the @{attribute sym} attribute, to be
wenzelm@30055
  1004
  used in fact expressions ``@{text "a [symmetric]"}'', or single-step
wenzelm@30055
  1005
  proofs ``@{command assume}~@{text "x = y"}~@{command then}~@{command
wenzelm@30055
  1006
  have}~@{text "y = x"}~@{command ".."}''.
wenzelm@30055
  1007
*}
wenzelm@30055
  1008
wenzelm@30068
  1009
end