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